src/HOL/Rings.thy
author haftmann
Mon, 04 Jul 2016 19:46:20 +0200
changeset 63375 59803048b0e8
parent 63359 99b51ba8da1c
child 63456 3365c8ec67bd
permissions -rw-r--r--
basic facts about almost everywhere fix bijections
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
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
    10
section \<open>Rings\<close>
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
62366
95c6cf433c91 more theorems
haftmann
parents: 62349
diff changeset
    13
imports Groups Set
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
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    21
text \<open>For the \<open>combine_numerals\<close> simproc\<close>
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    22
lemma combine_common_factor: "a * e + (b * e + c) = (a + b) * e + c"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    23
  by (simp add: distrib_right ac_simps)
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    24
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    25
end
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    26
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
    27
class mult_zero = times + zero +
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
    28
  assumes mult_zero_left [simp]: "0 * a = 0"
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
    29
  assumes mult_zero_right [simp]: "a * 0 = 0"
58195
1fee63e0377d added various facts
haftmann
parents: 57514
diff changeset
    30
begin
1fee63e0377d added various facts
haftmann
parents: 57514
diff changeset
    31
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    32
lemma mult_not_zero: "a * b \<noteq> 0 \<Longrightarrow> a \<noteq> 0 \<and> b \<noteq> 0"
58195
1fee63e0377d added various facts
haftmann
parents: 57514
diff changeset
    33
  by auto
1fee63e0377d added various facts
haftmann
parents: 57514
diff changeset
    34
1fee63e0377d added various facts
haftmann
parents: 57514
diff changeset
    35
end
21199
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
    36
58198
099ca49b5231 generalized
haftmann
parents: 58195
diff changeset
    37
class semiring_0 = semiring + comm_monoid_add + mult_zero
099ca49b5231 generalized
haftmann
parents: 58195
diff changeset
    38
29904
856f16a3b436 add class cancel_comm_monoid_add
huffman
parents: 29833
diff changeset
    39
class semiring_0_cancel = semiring + cancel_comm_monoid_add
25186
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
    40
begin
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    41
25186
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
    42
subclass semiring_0
28823
dcbef866c9e2 tuned unfold_locales invocation
haftmann
parents: 28559
diff changeset
    43
proof
21199
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
    44
  fix a :: 'a
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 44921
diff changeset
    45
  have "0 * a + 0 * a = 0 * a + 0" by (simp add: distrib_right [symmetric])
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    46
  then show "0 * a = 0" by (simp only: add_left_cancel)
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 44921
diff changeset
    47
  have "a * 0 + a * 0 = a * 0 + 0" by (simp add: distrib_left [symmetric])
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    48
  then show "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
    49
qed
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
    50
25186
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
    51
end
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    52
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
    53
class comm_semiring = ab_semigroup_add + ab_semigroup_mult +
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
    54
  assumes distrib: "(a + b) * c = a * c + b * c"
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    55
begin
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    56
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    57
subclass semiring
28823
dcbef866c9e2 tuned unfold_locales invocation
haftmann
parents: 28559
diff changeset
    58
proof
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    59
  fix a b c :: 'a
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    60
  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
    61
  have "a * (b + c) = (b + c) * a" by (simp add: ac_simps)
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    62
  also have "\<dots> = b * a + c * a" by (simp only: distrib)
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    63
  also have "\<dots> = 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
    64
  finally show "a * (b + c) = a * b + a * c" by blast
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    65
qed
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    66
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    67
end
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    68
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    69
class comm_semiring_0 = comm_semiring + comm_monoid_add + mult_zero
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    70
begin
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    71
27516
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
    72
subclass semiring_0 ..
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    73
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    74
end
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    75
29904
856f16a3b436 add class cancel_comm_monoid_add
huffman
parents: 29833
diff changeset
    76
class comm_semiring_0_cancel = comm_semiring + cancel_comm_monoid_add
25186
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
    77
begin
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
    78
27516
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
    79
subclass semiring_0_cancel ..
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
    80
28141
193c3ea0f63b instances comm_semiring_0_cancel < comm_semiring_0, comm_ring < comm_semiring_0_cancel
huffman
parents: 27651
diff changeset
    81
subclass comm_semiring_0 ..
193c3ea0f63b instances comm_semiring_0_cancel < comm_semiring_0, comm_ring < comm_semiring_0_cancel
huffman
parents: 27651
diff changeset
    82
25186
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
    83
end
21199
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
    84
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
    85
class zero_neq_one = zero + one +
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
    86
  assumes zero_neq_one [simp]: "0 \<noteq> 1"
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
    87
begin
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
    88
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
    89
lemma one_neq_zero [simp]: "1 \<noteq> 0"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    90
  by (rule not_sym) (rule zero_neq_one)
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
    91
54225
8a49a5a30284 generalized of_bool conversion
haftmann
parents: 54147
diff changeset
    92
definition of_bool :: "bool \<Rightarrow> 'a"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    93
  where "of_bool p = (if p then 1 else 0)"
54225
8a49a5a30284 generalized of_bool conversion
haftmann
parents: 54147
diff changeset
    94
8a49a5a30284 generalized of_bool conversion
haftmann
parents: 54147
diff changeset
    95
lemma of_bool_eq [simp, code]:
8a49a5a30284 generalized of_bool conversion
haftmann
parents: 54147
diff changeset
    96
  "of_bool False = 0"
8a49a5a30284 generalized of_bool conversion
haftmann
parents: 54147
diff changeset
    97
  "of_bool True = 1"
8a49a5a30284 generalized of_bool conversion
haftmann
parents: 54147
diff changeset
    98
  by (simp_all add: of_bool_def)
8a49a5a30284 generalized of_bool conversion
haftmann
parents: 54147
diff changeset
    99
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   100
lemma of_bool_eq_iff: "of_bool p = of_bool q \<longleftrightarrow> p = q"
54225
8a49a5a30284 generalized of_bool conversion
haftmann
parents: 54147
diff changeset
   101
  by (simp add: of_bool_def)
8a49a5a30284 generalized of_bool conversion
haftmann
parents: 54147
diff changeset
   102
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   103
lemma split_of_bool [split]: "P (of_bool p) \<longleftrightarrow> (p \<longrightarrow> P 1) \<and> (\<not> p \<longrightarrow> P 0)"
55187
6d0d93316daf split rules for of_bool, similar to if
haftmann
parents: 54703
diff changeset
   104
  by (cases p) simp_all
6d0d93316daf split rules for of_bool, similar to if
haftmann
parents: 54703
diff changeset
   105
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   106
lemma split_of_bool_asm: "P (of_bool p) \<longleftrightarrow> \<not> (p \<and> \<not> P 1 \<or> \<not> p \<and> \<not> P 0)"
55187
6d0d93316daf split rules for of_bool, similar to if
haftmann
parents: 54703
diff changeset
   107
  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
   108
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
   109
end
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   110
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   111
class semiring_1 = zero_neq_one + semiring_0 + monoid_mult
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   112
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
   113
text \<open>Abstract divisibility\<close>
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   114
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   115
class dvd = times
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   116
begin
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   117
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   118
definition dvd :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infix "dvd" 50)
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   119
  where "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
   120
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   121
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
   122
  unfolding dvd_def ..
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   123
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   124
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
   125
  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
   126
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   127
end
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   128
59009
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   129
context comm_monoid_mult
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   130
begin
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   131
59009
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   132
subclass dvd .
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   133
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   134
lemma dvd_refl [simp]: "a dvd a"
28559
55c003a5600a tuned default rules of (dvd)
haftmann
parents: 28141
diff changeset
   135
proof
55c003a5600a tuned default rules of (dvd)
haftmann
parents: 28141
diff changeset
   136
  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
   137
qed
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   138
62349
7c23469b5118 cleansed junk-producing interpretations for gcd/lcm on nat altogether
haftmann
parents: 62347
diff changeset
   139
lemma dvd_trans [trans]:
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   140
  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
   141
  shows "a dvd c"
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   142
proof -
28559
55c003a5600a tuned default rules of (dvd)
haftmann
parents: 28141
diff changeset
   143
  from assms obtain v where "b = a * v" by (auto elim!: dvdE)
55c003a5600a tuned default rules of (dvd)
haftmann
parents: 28141
diff changeset
   144
  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
   145
  ultimately have "c = a * (v * w)" by (simp add: mult.assoc)
28559
55c003a5600a tuned default rules of (dvd)
haftmann
parents: 28141
diff changeset
   146
  then show ?thesis ..
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   147
qed
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   148
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   149
lemma subset_divisors_dvd: "{c. c dvd a} \<subseteq> {c. c dvd b} \<longleftrightarrow> a dvd b"
62366
95c6cf433c91 more theorems
haftmann
parents: 62349
diff changeset
   150
  by (auto simp add: subset_iff intro: dvd_trans)
95c6cf433c91 more theorems
haftmann
parents: 62349
diff changeset
   151
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   152
lemma strict_subset_divisors_dvd: "{c. c dvd a} \<subset> {c. c dvd b} \<longleftrightarrow> a dvd b \<and> \<not> b dvd a"
62366
95c6cf433c91 more theorems
haftmann
parents: 62349
diff changeset
   153
  by (auto simp add: subset_iff intro: dvd_trans)
95c6cf433c91 more theorems
haftmann
parents: 62349
diff changeset
   154
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   155
lemma one_dvd [simp]: "1 dvd a"
59009
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   156
  by (auto intro!: dvdI)
28559
55c003a5600a tuned default rules of (dvd)
haftmann
parents: 28141
diff changeset
   157
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   158
lemma dvd_mult [simp]: "a dvd c \<Longrightarrow> a dvd (b * c)"
59009
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   159
  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
   160
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   161
lemma dvd_mult2 [simp]: "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
   162
  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
   163
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   164
lemma dvd_triv_right [simp]: "a dvd b * a"
59009
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   165
  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
   166
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   167
lemma dvd_triv_left [simp]: "a dvd a * b"
59009
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   168
  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
   169
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   170
lemma mult_dvd_mono:
30042
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29981
diff changeset
   171
  assumes "a dvd b"
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29981
diff changeset
   172
    and "c dvd d"
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   173
  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
   174
proof -
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
   175
  from \<open>a dvd b\<close> obtain b' where "b = a * b'" ..
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
   176
  moreover from \<open>c dvd d\<close> 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
   177
  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
   178
  then show ?thesis ..
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   179
qed
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   180
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   181
lemma dvd_mult_left: "a * b dvd c \<Longrightarrow> a dvd c"
59009
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   182
  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
   183
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   184
lemma dvd_mult_right: "a * b dvd c \<Longrightarrow> b dvd c"
59009
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   185
  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
   186
59009
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   187
end
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   188
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   189
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
   190
begin
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   191
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   192
subclass semiring_1 ..
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   193
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   194
lemma dvd_0_left_iff [simp]: "0 dvd a \<longleftrightarrow> a = 0"
59009
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   195
  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
   196
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   197
lemma dvd_0_right [iff]: "a dvd 0"
59009
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   198
proof
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   199
  show "0 = a * 0" by simp
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   200
qed
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   201
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   202
lemma dvd_0_left: "0 dvd a \<Longrightarrow> a = 0"
59009
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   203
  by simp
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   204
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   205
lemma dvd_add [simp]:
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   206
  assumes "a dvd b" and "a dvd c"
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   207
  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
   208
proof -
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
   209
  from \<open>a dvd b\<close> obtain b' where "b = a * b'" ..
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
   210
  moreover from \<open>a dvd c\<close> obtain c' where "c = a * c'" ..
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 44921
diff changeset
   211
  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
   212
  then show ?thesis ..
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   213
qed
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   214
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   215
end
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   216
29904
856f16a3b436 add class cancel_comm_monoid_add
huffman
parents: 29833
diff changeset
   217
class semiring_1_cancel = semiring + cancel_comm_monoid_add
856f16a3b436 add class cancel_comm_monoid_add
huffman
parents: 29833
diff changeset
   218
  + zero_neq_one + monoid_mult
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   219
begin
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
   220
27516
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
   221
subclass semiring_0_cancel ..
25512
4134f7c782e2 using intro_locales instead of unfold_locales if appropriate
haftmann
parents: 25450
diff changeset
   222
27516
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
   223
subclass semiring_1 ..
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   224
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   225
end
21199
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
   226
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   227
class comm_semiring_1_cancel =
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   228
  comm_semiring + cancel_comm_monoid_add + zero_neq_one + comm_monoid_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
   229
  assumes right_diff_distrib' [algebra_simps]: "a * (b - c) = a * b - a * c"
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   230
begin
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   231
27516
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
   232
subclass semiring_1_cancel ..
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
   233
subclass comm_semiring_0_cancel ..
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
   234
subclass comm_semiring_1 ..
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   235
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   236
lemma left_diff_distrib' [algebra_simps]: "(b - c) * a = b * a - c * a"
59816
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   237
  by (simp add: algebra_simps)
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   238
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   239
lemma dvd_add_times_triv_left_iff [simp]: "a dvd c * a + b \<longleftrightarrow> a dvd b"
59816
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   240
proof -
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   241
  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
   242
  proof
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   243
    assume ?Q
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   244
    then show ?P by simp
59816
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   245
  next
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   246
    assume ?P
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   247
    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
   248
    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
   249
    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
   250
    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
   251
    then show ?Q ..
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   252
  qed
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   253
  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
   254
qed
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   255
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   256
lemma dvd_add_times_triv_right_iff [simp]: "a dvd b + c * a \<longleftrightarrow> a dvd b"
59816
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   257
  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
   258
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   259
lemma dvd_add_triv_left_iff [simp]: "a dvd a + b \<longleftrightarrow> a dvd b"
59816
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   260
  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
   261
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   262
lemma dvd_add_triv_right_iff [simp]: "a dvd b + a \<longleftrightarrow> a dvd b"
59816
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   263
  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
   264
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   265
lemma dvd_add_right_iff:
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   266
  assumes "a dvd b"
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   267
  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
   268
proof
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   269
  assume ?P
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   270
  then obtain d where "b + c = a * d" ..
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
   271
  moreover from \<open>a dvd b\<close> obtain e where "b = a * e" ..
59816
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   272
  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
   273
  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
   274
  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
   275
  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
   276
  then show ?Q ..
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   277
next
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   278
  assume ?Q
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   279
  with assms show ?P by simp
59816
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   280
qed
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   281
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   282
lemma dvd_add_left_iff: "a dvd c \<Longrightarrow> a dvd b + c \<longleftrightarrow> a dvd b"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   283
  using dvd_add_right_iff [of a c b] by (simp add: ac_simps)
59816
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   284
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   285
end
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   286
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   287
class ring = semiring + ab_group_add
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   288
begin
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   289
27516
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
   290
subclass semiring_0_cancel ..
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   291
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
   292
text \<open>Distribution rules\<close>
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   293
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   294
lemma minus_mult_left: "- (a * b) = - a * b"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   295
  by (rule minus_unique) (simp add: distrib_right [symmetric])
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   296
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   297
lemma minus_mult_right: "- (a * b) = a * - b"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   298
  by (rule minus_unique) (simp add: distrib_left [symmetric])
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   299
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   300
text \<open>Extract signs from products\<close>
54147
97a8ff4e4ac9 killed most "no_atp", to make Sledgehammer more complete
blanchet
parents: 52435
diff changeset
   301
lemmas mult_minus_left [simp] = minus_mult_left [symmetric]
97a8ff4e4ac9 killed most "no_atp", to make Sledgehammer more complete
blanchet
parents: 52435
diff changeset
   302
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
   303
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   304
lemma minus_mult_minus [simp]: "- a * - b = a * b"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   305
  by simp
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   306
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   307
lemma minus_mult_commute: "- a * b = a * - b"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   308
  by simp
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
   309
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   310
lemma right_diff_distrib [algebra_simps]: "a * (b - c) = a * b - a * c"
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 54225
diff changeset
   311
  using distrib_left [of a b "-c "] by simp
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
   312
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   313
lemma left_diff_distrib [algebra_simps]: "(a - b) * c = a * c - b * c"
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 54225
diff changeset
   314
  using distrib_right [of a "- b" c] by simp
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   315
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   316
lemmas ring_distribs = distrib_left distrib_right left_diff_distrib right_diff_distrib
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   317
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   318
lemma eq_add_iff1: "a * e + c = b * e + d \<longleftrightarrow> (a - b) * e + c = d"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   319
  by (simp add: algebra_simps)
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   320
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   321
lemma eq_add_iff2: "a * e + c = b * e + d \<longleftrightarrow> c = (b - a) * e + d"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   322
  by (simp add: algebra_simps)
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   323
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   324
end
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   325
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   326
lemmas ring_distribs = distrib_left distrib_right left_diff_distrib right_diff_distrib
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   327
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   328
class comm_ring = comm_semiring + ab_group_add
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   329
begin
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   330
27516
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
   331
subclass ring ..
28141
193c3ea0f63b instances comm_semiring_0_cancel < comm_semiring_0, comm_ring < comm_semiring_0_cancel
huffman
parents: 27651
diff changeset
   332
subclass comm_semiring_0_cancel ..
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   333
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   334
lemma square_diff_square_factored: "x * x - y * y = (x + y) * (x - y)"
44350
63cddfbc5a09 replace lemma realpow_two_diff with new lemma square_diff_square_factored
huffman
parents: 44346
diff changeset
   335
  by (simp add: algebra_simps)
63cddfbc5a09 replace lemma realpow_two_diff with new lemma square_diff_square_factored
huffman
parents: 44346
diff changeset
   336
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   337
end
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   338
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   339
class ring_1 = ring + zero_neq_one + monoid_mult
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   340
begin
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   341
27516
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
   342
subclass semiring_1_cancel ..
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   343
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   344
lemma square_diff_one_factored: "x * x - 1 = (x + 1) * (x - 1)"
44346
00dd3c4dabe0 rename real_squared_diff_one_factored to square_diff_one_factored and move to Rings.thy
huffman
parents: 44064
diff changeset
   345
  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
   346
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   347
end
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   348
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   349
class comm_ring_1 = comm_ring + zero_neq_one + comm_monoid_mult
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   350
begin
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   351
27516
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
   352
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
   353
subclass comm_semiring_1_cancel
59816
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   354
  by unfold_locales (simp add: algebra_simps)
58647
fce800afeec7 more facts about abstract divisibility
haftmann
parents: 58198
diff changeset
   355
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
   356
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
   357
proof
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   358
  assume "x dvd - y"
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   359
  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
   360
  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
   361
next
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   362
  assume "x dvd y"
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   363
  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
   364
  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
   365
qed
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   366
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
   367
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
   368
proof
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   369
  assume "- x dvd y"
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   370
  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
   371
  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
   372
  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
   373
next
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   374
  assume "x dvd y"
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   375
  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
   376
  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
   377
  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
   378
qed
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   379
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   380
lemma dvd_diff [simp]: "x dvd y \<Longrightarrow> x dvd z \<Longrightarrow> x dvd (y - z)"
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 54225
diff changeset
   381
  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
   382
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   383
end
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   384
59833
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   385
class semiring_no_zero_divisors = semiring_0 +
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   386
  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
   387
begin
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   388
59833
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   389
lemma divisors_zero:
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   390
  assumes "a * b = 0"
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   391
  shows "a = 0 \<or> b = 0"
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   392
proof (rule classical)
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   393
  assume "\<not> ?thesis"
59833
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   394
  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
   395
  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
   396
  with assms show ?thesis by simp
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   397
qed
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   398
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   399
lemma mult_eq_0_iff [simp]: "a * b = 0 \<longleftrightarrow> a = 0 \<or> b = 0"
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   400
proof (cases "a = 0 \<or> b = 0")
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   401
  case False
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   402
  then have "a \<noteq> 0" and "b \<noteq> 0" by auto
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   403
    then show ?thesis using no_zero_divisors by simp
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   404
next
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   405
  case True
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   406
  then show ?thesis by auto
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   407
qed
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   408
58952
5d82cdef6c1b equivalence rules for structures without zero divisors
haftmann
parents: 58889
diff changeset
   409
end
5d82cdef6c1b equivalence rules for structures without zero divisors
haftmann
parents: 58889
diff changeset
   410
62481
b5d8e57826df tuned bootstrap order to provide type classes in a more sensible order
haftmann
parents: 62390
diff changeset
   411
class semiring_1_no_zero_divisors = semiring_1 + semiring_no_zero_divisors
b5d8e57826df tuned bootstrap order to provide type classes in a more sensible order
haftmann
parents: 62390
diff changeset
   412
60516
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   413
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
   414
  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
   415
    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
   416
begin
5d82cdef6c1b equivalence rules for structures without zero divisors
haftmann
parents: 58889
diff changeset
   417
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   418
lemma mult_left_cancel: "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
   419
  by simp
56217
dc429a5b13c4 Some rationalisation of basic lemmas
paulson <lp15@cam.ac.uk>
parents: 55912
diff changeset
   420
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   421
lemma mult_right_cancel: "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
   422
  by simp
56217
dc429a5b13c4 Some rationalisation of basic lemmas
paulson <lp15@cam.ac.uk>
parents: 55912
diff changeset
   423
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   424
end
22990
775e9de3db48 added classes ring_no_zero_divisors and dom (non-commutative version of idom);
huffman
parents: 22987
diff changeset
   425
60516
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   426
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
   427
begin
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   428
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   429
subclass semiring_no_zero_divisors_cancel
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   430
proof
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   431
  fix a b c
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   432
  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
   433
    by (simp add: algebra_simps)
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   434
  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
   435
    by auto
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   436
  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
   437
  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
   438
    by (simp add: algebra_simps)
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   439
  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
   440
    by auto
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   441
  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
   442
qed
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   443
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   444
end
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   445
23544
4b4165cb3e0d rename class dom to ring_1_no_zero_divisors
huffman
parents: 23527
diff changeset
   446
class ring_1_no_zero_divisors = ring_1 + ring_no_zero_divisors
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
   447
begin
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
   448
62481
b5d8e57826df tuned bootstrap order to provide type classes in a more sensible order
haftmann
parents: 62390
diff changeset
   449
subclass semiring_1_no_zero_divisors ..
b5d8e57826df tuned bootstrap order to provide type classes in a more sensible order
haftmann
parents: 62390
diff changeset
   450
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   451
lemma square_eq_1_iff: "x * x = 1 \<longleftrightarrow> x = 1 \<or> x = - 1"
36821
9207505d1ee5 move lemma real_mult_is_one to Rings.thy, renamed to square_eq_1_iff
huffman
parents: 36719
diff changeset
   452
proof -
9207505d1ee5 move lemma real_mult_is_one to Rings.thy, renamed to square_eq_1_iff
huffman
parents: 36719
diff changeset
   453
  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
   454
    by (simp add: algebra_simps)
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   455
  then have "x * x = 1 \<longleftrightarrow> (x - 1) * (x + 1) = 0"
36821
9207505d1ee5 move lemma real_mult_is_one to Rings.thy, renamed to square_eq_1_iff
huffman
parents: 36719
diff changeset
   456
    by simp
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   457
  then show ?thesis
36821
9207505d1ee5 move lemma real_mult_is_one to Rings.thy, renamed to square_eq_1_iff
huffman
parents: 36719
diff changeset
   458
    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
   459
qed
9207505d1ee5 move lemma real_mult_is_one to Rings.thy, renamed to square_eq_1_iff
huffman
parents: 36719
diff changeset
   460
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   461
lemma mult_cancel_right1 [simp]: "c = b * c \<longleftrightarrow> c = 0 \<or> b = 1"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   462
  using mult_cancel_right [of 1 c b] by auto
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
   463
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   464
lemma mult_cancel_right2 [simp]: "a * c = c \<longleftrightarrow> c = 0 \<or> a = 1"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   465
  using mult_cancel_right [of a c 1] 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
   466
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   467
lemma mult_cancel_left1 [simp]: "c = c * b \<longleftrightarrow> c = 0 \<or> b = 1"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   468
  using mult_cancel_left [of c 1 b] by force
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
   469
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   470
lemma mult_cancel_left2 [simp]: "c * a = c \<longleftrightarrow> c = 0 \<or> a = 1"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   471
  using mult_cancel_left [of c a 1] by simp
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
   472
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
   473
end
22990
775e9de3db48 added classes ring_no_zero_divisors and dom (non-commutative version of idom);
huffman
parents: 22987
diff changeset
   474
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
   475
class semidom = comm_semiring_1_cancel + semiring_no_zero_divisors
62481
b5d8e57826df tuned bootstrap order to provide type classes in a more sensible order
haftmann
parents: 62390
diff changeset
   476
begin
b5d8e57826df tuned bootstrap order to provide type classes in a more sensible order
haftmann
parents: 62390
diff changeset
   477
b5d8e57826df tuned bootstrap order to provide type classes in a more sensible order
haftmann
parents: 62390
diff changeset
   478
subclass semiring_1_no_zero_divisors ..
b5d8e57826df tuned bootstrap order to provide type classes in a more sensible order
haftmann
parents: 62390
diff changeset
   479
b5d8e57826df tuned bootstrap order to provide type classes in a more sensible order
haftmann
parents: 62390
diff changeset
   480
end
59833
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   481
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   482
class idom = comm_ring_1 + semiring_no_zero_divisors
25186
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
   483
begin
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   484
59833
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   485
subclass semidom ..
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   486
27516
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
   487
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
   488
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   489
lemma dvd_mult_cancel_right [simp]: "a * c dvd b * c \<longleftrightarrow> c = 0 \<or> a dvd b"
29981
7d0ed261b712 generalize int_dvd_cancel_factor simproc to idom class
huffman
parents: 29949
diff changeset
   490
proof -
7d0ed261b712 generalize int_dvd_cancel_factor simproc to idom class
huffman
parents: 29949
diff changeset
   491
  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
   492
    unfolding dvd_def by (simp add: ac_simps)
29981
7d0ed261b712 generalize int_dvd_cancel_factor simproc to idom class
huffman
parents: 29949
diff changeset
   493
  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
   494
    unfolding dvd_def by simp
7d0ed261b712 generalize int_dvd_cancel_factor simproc to idom class
huffman
parents: 29949
diff changeset
   495
  finally show ?thesis .
7d0ed261b712 generalize int_dvd_cancel_factor simproc to idom class
huffman
parents: 29949
diff changeset
   496
qed
7d0ed261b712 generalize int_dvd_cancel_factor simproc to idom class
huffman
parents: 29949
diff changeset
   497
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   498
lemma dvd_mult_cancel_left [simp]: "c * a dvd c * b \<longleftrightarrow> c = 0 \<or> a dvd b"
29981
7d0ed261b712 generalize int_dvd_cancel_factor simproc to idom class
huffman
parents: 29949
diff changeset
   499
proof -
7d0ed261b712 generalize int_dvd_cancel_factor simproc to idom class
huffman
parents: 29949
diff changeset
   500
  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
   501
    unfolding dvd_def by (simp add: ac_simps)
29981
7d0ed261b712 generalize int_dvd_cancel_factor simproc to idom class
huffman
parents: 29949
diff changeset
   502
  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
   503
    unfolding dvd_def by simp
7d0ed261b712 generalize int_dvd_cancel_factor simproc to idom class
huffman
parents: 29949
diff changeset
   504
  finally show ?thesis .
7d0ed261b712 generalize int_dvd_cancel_factor simproc to idom class
huffman
parents: 29949
diff changeset
   505
qed
7d0ed261b712 generalize int_dvd_cancel_factor simproc to idom class
huffman
parents: 29949
diff changeset
   506
60516
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   507
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
   508
proof
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   509
  assume "a * a = b * b"
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   510
  then have "(a - b) * (a + b) = 0"
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   511
    by (simp add: algebra_simps)
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   512
  then show "a = b \<or> a = - b"
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   513
    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
   514
next
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   515
  assume "a = b \<or> a = - b"
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   516
  then show "a * a = b * b" by auto
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   517
qed
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   518
25186
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
   519
end
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   520
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
   521
text \<open>
35302
4bc6b4d70e08 tuned text
haftmann
parents: 35216
diff changeset
   522
  The theory of partially ordered rings is taken from the books:
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   523
    \<^item> \<^emph>\<open>Lattice Theory\<close> by Garret Birkhoff, American Mathematical Society, 1979
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   524
    \<^item> \<^emph>\<open>Partially Ordered Algebraic Systems\<close>, Pergamon Press, 1963
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   525
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
   526
  Most of the used notions can also be looked up in
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   527
    \<^item> @{url "http://www.mathworld.com"} by Eric Weisstein et. al.
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   528
    \<^item> \<^emph>\<open>Algebra I\<close> by van der Waerden, Springer
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
   529
\<close>
35302
4bc6b4d70e08 tuned text
haftmann
parents: 35216
diff changeset
   530
60353
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   531
class divide =
60429
d3d1e185cd63 uniform _ div _ as infix syntax for ring division
haftmann
parents: 60353
diff changeset
   532
  fixes divide :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  (infixl "div" 70)
60353
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   533
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
   534
setup \<open>Sign.add_const_constraint (@{const_name "divide"}, SOME @{typ "'a \<Rightarrow> 'a \<Rightarrow> 'a"})\<close>
60353
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   535
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   536
context semiring
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   537
begin
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   538
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   539
lemma [field_simps]:
60429
d3d1e185cd63 uniform _ div _ as infix syntax for ring division
haftmann
parents: 60353
diff changeset
   540
  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
   541
    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
   542
  by (rule distrib_left distrib_right)+
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   543
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   544
end
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   545
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   546
context ring
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   547
begin
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   548
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   549
lemma [field_simps]:
60429
d3d1e185cd63 uniform _ div _ as infix syntax for ring division
haftmann
parents: 60353
diff changeset
   550
  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
   551
    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
   552
  by (rule left_diff_distrib right_diff_distrib)+
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   553
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   554
end
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   555
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
   556
setup \<open>Sign.add_const_constraint (@{const_name "divide"}, SOME @{typ "'a::divide \<Rightarrow> 'a \<Rightarrow> 'a"})\<close>
60353
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   557
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   558
class semidom_divide = semidom + divide +
60429
d3d1e185cd63 uniform _ div _ as infix syntax for ring division
haftmann
parents: 60353
diff changeset
   559
  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
   560
  assumes divide_zero [simp]: "a div 0 = 0"
60353
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   561
begin
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   562
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   563
lemma nonzero_mult_divide_cancel_left [simp]: "a \<noteq> 0 \<Longrightarrow> (a * b) div a = b"
60353
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   564
  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
   565
60516
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   566
subclass semiring_no_zero_divisors_cancel
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   567
proof
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   568
  show *: "a * c = b * c \<longleftrightarrow> c = 0 \<or> a = b" for a b c
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   569
  proof (cases "c = 0")
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   570
    case True
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   571
    then show ?thesis by simp
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   572
  next
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   573
    case False
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   574
    {
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   575
      assume "a * c = b * c"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   576
      then have "a * c div c = b * c div c"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   577
        by simp
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   578
      with False have "a = b"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   579
        by simp
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   580
    }
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   581
    then show ?thesis by auto
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   582
  qed
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   583
  show "c * a = c * b \<longleftrightarrow> c = 0 \<or> a = b" for a b c
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   584
    using * [of a c b] by (simp add: ac_simps)
60516
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   585
qed
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   586
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   587
lemma div_self [simp]: "a \<noteq> 0 \<Longrightarrow> a div a = 1"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   588
  using nonzero_mult_divide_cancel_left [of a 1] by simp
60516
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   589
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   590
lemma divide_zero_left [simp]: "0 div a = 0"
60570
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   591
proof (cases "a = 0")
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   592
  case True
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   593
  then show ?thesis by simp
60570
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   594
next
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   595
  case False
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   596
  then have "a * 0 div a = 0"
60570
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   597
    by (rule nonzero_mult_divide_cancel_left)
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   598
  then show ?thesis by simp
62376
85f38d5f8807 Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents: 62366
diff changeset
   599
qed
60570
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   600
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   601
lemma divide_1 [simp]: "a div 1 = a"
60690
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   602
  using nonzero_mult_divide_cancel_left [of 1 a] by simp
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   603
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   604
end
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   605
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   606
class idom_divide = idom + semidom_divide
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   607
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   608
class algebraic_semidom = semidom_divide
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   609
begin
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   610
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   611
text \<open>
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   612
  Class @{class algebraic_semidom} enriches a integral domain
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   613
  by notions from algebra, like units in a ring.
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   614
  It is a separate class to avoid spoiling fields with notions
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   615
  which are degenerated there.
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   616
\<close>
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   617
60690
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   618
lemma dvd_times_left_cancel_iff [simp]:
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   619
  assumes "a \<noteq> 0"
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   620
  shows "a * b dvd a * c \<longleftrightarrow> b dvd c" (is "?P \<longleftrightarrow> ?Q")
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   621
proof
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   622
  assume ?P
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   623
  then obtain d where "a * c = a * b * d" ..
60690
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   624
  with assms have "c = b * d" by (simp add: ac_simps)
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   625
  then show ?Q ..
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   626
next
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   627
  assume ?Q
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   628
  then obtain d where "c = b * d" ..
60690
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   629
  then have "a * c = a * b * d" by (simp add: ac_simps)
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   630
  then show ?P ..
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   631
qed
62376
85f38d5f8807 Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents: 62366
diff changeset
   632
60690
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   633
lemma dvd_times_right_cancel_iff [simp]:
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   634
  assumes "a \<noteq> 0"
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   635
  shows "b * a dvd c * a \<longleftrightarrow> b dvd c" (is "?P \<longleftrightarrow> ?Q")
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   636
  using dvd_times_left_cancel_iff [of a b c] assms by (simp add: ac_simps)
62376
85f38d5f8807 Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents: 62366
diff changeset
   637
60690
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   638
lemma div_dvd_iff_mult:
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   639
  assumes "b \<noteq> 0" and "b dvd a"
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   640
  shows "a div b dvd c \<longleftrightarrow> a dvd c * b"
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   641
proof -
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   642
  from \<open>b dvd a\<close> obtain d where "a = b * d" ..
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   643
  with \<open>b \<noteq> 0\<close> show ?thesis by (simp add: ac_simps)
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   644
qed
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   645
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   646
lemma dvd_div_iff_mult:
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   647
  assumes "c \<noteq> 0" and "c dvd b"
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   648
  shows "a dvd b div c \<longleftrightarrow> a * c dvd b"
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   649
proof -
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   650
  from \<open>c dvd b\<close> obtain d where "b = c * d" ..
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   651
  with \<open>c \<noteq> 0\<close> show ?thesis by (simp add: mult.commute [of a])
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   652
qed
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   653
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   654
lemma div_dvd_div [simp]:
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   655
  assumes "a dvd b" and "a dvd c"
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   656
  shows "b div a dvd c div a \<longleftrightarrow> b dvd c"
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   657
proof (cases "a = 0")
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   658
  case True
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   659
  with assms show ?thesis by simp
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   660
next
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   661
  case False
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   662
  moreover from assms obtain k l where "b = a * k" and "c = a * l"
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   663
    by (auto elim!: dvdE)
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   664
  ultimately show ?thesis by simp
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   665
qed
60353
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   666
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   667
lemma div_add [simp]:
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   668
  assumes "c dvd a" and "c dvd b"
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   669
  shows "(a + b) div c = a div c + b div c"
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   670
proof (cases "c = 0")
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   671
  case True
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   672
  then show ?thesis by simp
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   673
next
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   674
  case False
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   675
  moreover from assms obtain k l where "a = c * k" and "b = c * l"
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   676
    by (auto elim!: dvdE)
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   677
  moreover have "c * k + c * l = c * (k + l)"
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   678
    by (simp add: algebra_simps)
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   679
  ultimately show ?thesis
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   680
    by simp
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   681
qed
60517
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   682
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   683
lemma div_mult_div_if_dvd:
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   684
  assumes "b dvd a" and "d dvd c"
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   685
  shows "(a div b) * (c div d) = (a * c) div (b * d)"
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   686
proof (cases "b = 0 \<or> c = 0")
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   687
  case True
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   688
  with assms show ?thesis by auto
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   689
next
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   690
  case False
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   691
  moreover from assms obtain k l where "a = b * k" and "c = d * l"
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   692
    by (auto elim!: dvdE)
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   693
  moreover have "b * k * (d * l) div (b * d) = (b * d) * (k * l) div (b * d)"
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   694
    by (simp add: ac_simps)
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   695
  ultimately show ?thesis by simp
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   696
qed
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   697
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   698
lemma dvd_div_eq_mult:
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   699
  assumes "a \<noteq> 0" and "a dvd b"
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   700
  shows "b div a = c \<longleftrightarrow> b = c * a"
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   701
proof
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   702
  assume "b = c * a"
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   703
  then show "b div a = c" by (simp add: assms)
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   704
next
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   705
  assume "b div a = c"
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   706
  then have "b div a * a = c * a" by simp
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   707
  moreover from assms have "b div a * a = b"
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   708
    by (auto elim!: dvdE simp add: ac_simps)
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   709
  ultimately show "b = c * a" by simp
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   710
qed
60688
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
   711
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   712
lemma dvd_div_mult_self [simp]: "a dvd b \<Longrightarrow> b div a * a = b"
60517
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   713
  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
   714
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   715
lemma dvd_mult_div_cancel [simp]: "a dvd b \<Longrightarrow> a * (b div a) = b"
60517
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   716
  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
   717
60517
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   718
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
   719
  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
   720
  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
   721
proof (cases "c = 0")
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   722
  case True
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   723
  then show ?thesis by simp
60517
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   724
next
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   725
  case False
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   726
  from assms obtain d where "b = c * d" ..
60517
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   727
  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
   728
    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
   729
  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
   730
qed
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   731
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   732
lemma dvd_div_mult: "c dvd b \<Longrightarrow> b div c * a = (b * a) div c"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   733
  using div_mult_swap [of c b a] by (simp add: ac_simps)
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
60570
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   735
lemma dvd_div_mult2_eq:
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   736
  assumes "b * c dvd a"
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   737
  shows "a div (b * c) = a div b div c"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   738
proof -
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   739
  from assms obtain k where "a = b * c * k" ..
60570
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   740
  then show ?thesis
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   741
    by (cases "b = 0 \<or> c = 0") (auto, simp add: ac_simps)
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   742
qed
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   743
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   744
lemma dvd_div_div_eq_mult:
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   745
  assumes "a \<noteq> 0" "c \<noteq> 0" and "a dvd b" "c dvd d"
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   746
  shows "b div a = d div c \<longleftrightarrow> b * c = a * d" (is "?P \<longleftrightarrow> ?Q")
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   747
proof -
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   748
  from assms have "a * c \<noteq> 0" by simp
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   749
  then have "?P \<longleftrightarrow> b div a * (a * c) = d div c * (a * c)"
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   750
    by simp
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   751
  also have "\<dots> \<longleftrightarrow> (a * (b div a)) * c = (c * (d div c)) * a"
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   752
    by (simp add: ac_simps)
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   753
  also have "\<dots> \<longleftrightarrow> (a * b div a) * c = (c * d div c) * a"
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   754
    using assms by (simp add: div_mult_swap)
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   755
  also have "\<dots> \<longleftrightarrow> ?Q"
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   756
    using assms by (simp add: ac_simps)
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   757
  finally show ?thesis .
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   758
qed
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   759
63359
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63325
diff changeset
   760
lemma dvd_mult_imp_div:
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63325
diff changeset
   761
  assumes "a * c dvd b"
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63325
diff changeset
   762
  shows "a dvd b div c"
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63325
diff changeset
   763
proof (cases "c = 0")
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63325
diff changeset
   764
  case True then show ?thesis by simp
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63325
diff changeset
   765
next
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63325
diff changeset
   766
  case False
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63325
diff changeset
   767
  from \<open>a * c dvd b\<close> obtain d where "b = a * c * d" ..
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63325
diff changeset
   768
  with False show ?thesis by (simp add: mult.commute [of a] mult.assoc)
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63325
diff changeset
   769
qed
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63325
diff changeset
   770
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
   771
60517
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   772
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
   773
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   774
abbreviation is_unit :: "'a \<Rightarrow> bool"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   775
  where "is_unit a \<equiv> a dvd 1"
60517
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   776
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   777
lemma not_is_unit_0 [simp]: "\<not> is_unit 0"
60517
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   778
  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
   779
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   780
lemma unit_imp_dvd [dest]: "is_unit b \<Longrightarrow> b dvd a"
60517
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 (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
   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_dvdE:
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   784
  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
   785
  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
   786
proof -
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   787
  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
   788
  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
   789
  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
   790
  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
   791
qed
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   792
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   793
lemma dvd_unit_imp_unit: "a dvd b \<Longrightarrow> is_unit b \<Longrightarrow> is_unit a"
60517
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   794
  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
   795
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   796
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
   797
  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
   798
  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
   799
proof -
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   800
  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
   801
  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
   802
qed
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   803
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   804
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
   805
  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
   806
  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
   807
    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
   808
    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
   809
proof (rule that)
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62626
diff changeset
   810
  define b where "b = 1 div a"
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
  then show "1 div a = b" by simp
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   812
  from assms b_def show "is_unit b" by simp
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   813
  with assms show "a \<noteq> 0" and "b \<noteq> 0" by auto
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   814
  from assms b_def show "a * b = 1" by simp
60517
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   815
  then have "1 = a * b" ..
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
   816
  with b_def \<open>b \<noteq> 0\<close> show "1 div b = a" by simp
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   817
  from assms have "a dvd c" ..
60517
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   818
  then obtain d where "c = a * d" ..
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
   819
  with \<open>a \<noteq> 0\<close> \<open>a * b = 1\<close> show "c div a = c * b"
60517
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   820
    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
   821
qed
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   822
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   823
lemma unit_prod [intro]: "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
   824
  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
   825
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   826
lemma is_unit_mult_iff: "is_unit (a * b) \<longleftrightarrow> is_unit a \<and> is_unit b"
62366
95c6cf433c91 more theorems
haftmann
parents: 62349
diff changeset
   827
  by (auto dest: dvd_mult_left dvd_mult_right)
95c6cf433c91 more theorems
haftmann
parents: 62349
diff changeset
   828
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   829
lemma unit_div [intro]: "is_unit a \<Longrightarrow> is_unit b \<Longrightarrow> is_unit (a div b)"
60517
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   830
  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
   831
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   832
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
   833
  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
   834
  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
   835
proof
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   836
  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
   837
  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
   838
    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
   839
next
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   840
  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
   841
  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
   842
  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
   843
    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
   844
  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
   845
qed
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   846
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   847
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
   848
  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
   849
  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
   850
proof
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   851
  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
   852
  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
   853
    by (subst mult_assoc [symmetric]) simp
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   854
  also from assms have "b * (1 div b) = 1"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   855
    by (rule is_unitE) simp
60517
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   856
  finally have "c * b dvd c" by simp
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
   857
  with \<open>a dvd c * b\<close> show "a dvd c" by (rule dvd_trans)
60517
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   858
next
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   859
  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
   860
  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
   861
qed
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   862
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   863
lemma div_unit_dvd_iff: "is_unit b \<Longrightarrow> a div b dvd c \<longleftrightarrow> a dvd c"
60517
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   864
  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
   865
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   866
lemma dvd_div_unit_iff: "is_unit b \<Longrightarrow> a dvd c div b \<longleftrightarrow> a dvd c"
60517
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 (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
   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
lemmas unit_dvd_iff = mult_unit_dvd_iff div_unit_dvd_iff
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   870
  dvd_mult_unit_iff dvd_div_unit_iff  (* FIXME consider named_theorems *)
60517
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   871
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   872
lemma unit_mult_div_div [simp]: "is_unit a \<Longrightarrow> b * (1 div a) = b div a"
60517
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   873
  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
   874
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   875
lemma unit_div_mult_self [simp]: "is_unit a \<Longrightarrow> b div a * a = b"
60517
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   876
  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
   877
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   878
lemma unit_div_1_div_1 [simp]: "is_unit a \<Longrightarrow> 1 div (1 div a) = a"
60517
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   879
  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
   880
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   881
lemma unit_div_mult_swap: "is_unit c \<Longrightarrow> a * (b div c) = (a * b) div c"
60517
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   882
  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
   883
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   884
lemma unit_div_commute: "is_unit b \<Longrightarrow> (a div b) * c = (a * c) div b"
60517
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   885
  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
   886
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   887
lemma unit_eq_div1: "is_unit b \<Longrightarrow> a div b = c \<longleftrightarrow> a = c * b"
60517
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   888
  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
   889
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   890
lemma unit_eq_div2: "is_unit b \<Longrightarrow> a = c div b \<longleftrightarrow> a * b = c"
60517
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   891
  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
   892
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   893
lemma unit_mult_left_cancel: "is_unit a \<Longrightarrow> a * b = a * c \<longleftrightarrow> b = c"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   894
  using 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
   895
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   896
lemma unit_mult_right_cancel: "is_unit a \<Longrightarrow> b * a = c * a \<longleftrightarrow> b = c"
60517
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   897
  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
   898
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   899
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
   900
  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
   901
  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
   902
proof -
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   903
  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
   904
  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
   905
    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
   906
  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
   907
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
   908
60570
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   909
lemma is_unit_div_mult2_eq:
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   910
  assumes "is_unit b" and "is_unit c"
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   911
  shows "a div (b * c) = a div b div c"
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   912
proof -
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   913
  from assms have "is_unit (b * c)"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   914
    by (simp add: unit_prod)
60570
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   915
  then have "b * c dvd a"
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   916
    by (rule unit_imp_dvd)
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   917
  then show ?thesis
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   918
    by (rule dvd_div_mult2_eq)
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   919
qed
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   920
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
   921
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
   922
  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
   923
  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
   924
  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
   925
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   926
lemma is_unit_divide_mult_cancel_left:
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   927
  assumes "a \<noteq> 0" and "is_unit b"
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   928
  shows "a div (a * b) = 1 div b"
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   929
proof -
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   930
  from assms have "a div (a * b) = a div a div b"
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   931
    by (simp add: mult_unit_dvd_iff dvd_div_mult2_eq)
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   932
  with assms show ?thesis by simp
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   933
qed
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   934
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   935
lemma is_unit_divide_mult_cancel_right:
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   936
  assumes "a \<noteq> 0" and "is_unit b"
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   937
  shows "a div (b * a) = 1 div b"
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   938
  using assms is_unit_divide_mult_cancel_left [of a b] by (simp add: ac_simps)
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   939
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   940
end
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   941
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   942
class normalization_semidom = algebraic_semidom +
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   943
  fixes normalize :: "'a \<Rightarrow> 'a"
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   944
    and unit_factor :: "'a \<Rightarrow> 'a"
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   945
  assumes unit_factor_mult_normalize [simp]: "unit_factor a * normalize a = a"
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   946
  assumes normalize_0 [simp]: "normalize 0 = 0"
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   947
    and unit_factor_0 [simp]: "unit_factor 0 = 0"
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   948
  assumes is_unit_normalize:
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   949
    "is_unit a  \<Longrightarrow> normalize a = 1"
62376
85f38d5f8807 Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents: 62366
diff changeset
   950
  assumes unit_factor_is_unit [iff]:
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   951
    "a \<noteq> 0 \<Longrightarrow> is_unit (unit_factor a)"
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   952
  assumes unit_factor_mult: "unit_factor (a * b) = unit_factor a * unit_factor b"
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   953
begin
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   954
60688
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
   955
text \<open>
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
   956
  Class @{class normalization_semidom} cultivates the idea that
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
   957
  each integral domain can be split into equivalence classes
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
   958
  whose representants are associated, i.e. divide each other.
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
   959
  @{const normalize} specifies a canonical representant for each equivalence
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
   960
  class.  The rationale behind this is that it is easier to reason about equality
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
   961
  than equivalences, hence we prefer to think about equality of normalized
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
   962
  values rather than associated elements.
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
   963
\<close>
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
   964
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   965
lemma unit_factor_dvd [simp]: "a \<noteq> 0 \<Longrightarrow> unit_factor a dvd b"
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   966
  by (rule unit_imp_dvd) simp
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   967
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   968
lemma unit_factor_self [simp]: "unit_factor a dvd a"
62376
85f38d5f8807 Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents: 62366
diff changeset
   969
  by (cases "a = 0") simp_all
85f38d5f8807 Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents: 62366
diff changeset
   970
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   971
lemma normalize_mult_unit_factor [simp]: "normalize a * unit_factor a = a"
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   972
  using unit_factor_mult_normalize [of a] by (simp add: ac_simps)
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   973
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   974
lemma normalize_eq_0_iff [simp]: "normalize a = 0 \<longleftrightarrow> a = 0"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   975
  (is "?P \<longleftrightarrow> ?Q")
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   976
proof
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   977
  assume ?P
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   978
  moreover have "unit_factor a * normalize a = a" by simp
62376
85f38d5f8807 Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents: 62366
diff changeset
   979
  ultimately show ?Q by simp
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   980
next
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   981
  assume ?Q
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   982
  then show ?P by simp
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   983
qed
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   984
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   985
lemma unit_factor_eq_0_iff [simp]: "unit_factor a = 0 \<longleftrightarrow> a = 0"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   986
  (is "?P \<longleftrightarrow> ?Q")
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   987
proof
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   988
  assume ?P
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   989
  moreover have "unit_factor a * normalize a = a" by simp
62376
85f38d5f8807 Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents: 62366
diff changeset
   990
  ultimately show ?Q by simp
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   991
next
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   992
  assume ?Q
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   993
  then show ?P by simp
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   994
qed
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   995
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
   996
lemma is_unit_unit_factor:
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   997
  assumes "is_unit a"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   998
  shows "unit_factor a = a"
62376
85f38d5f8807 Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents: 62366
diff changeset
   999
proof -
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1000
  from assms have "normalize a = 1" by (rule is_unit_normalize)
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1001
  moreover from unit_factor_mult_normalize have "unit_factor a * normalize a = a" .
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1002
  ultimately show ?thesis by simp
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1003
qed
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1004
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1005
lemma unit_factor_1 [simp]: "unit_factor 1 = 1"
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1006
  by (rule is_unit_unit_factor) simp
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1007
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1008
lemma normalize_1 [simp]: "normalize 1 = 1"
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1009
  by (rule is_unit_normalize) simp
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1010
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1011
lemma normalize_1_iff: "normalize a = 1 \<longleftrightarrow> is_unit a"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1012
  (is "?P \<longleftrightarrow> ?Q")
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1013
proof
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1014
  assume ?Q
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1015
  then show ?P by (rule is_unit_normalize)
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1016
next
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1017
  assume ?P
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1018
  then have "a \<noteq> 0" by auto
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1019
  from \<open>?P\<close> have "unit_factor a * normalize a = unit_factor a * 1"
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1020
    by simp
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1021
  then have "unit_factor a = a"
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1022
    by simp
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1023
  moreover have "is_unit (unit_factor a)"
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1024
    using \<open>a \<noteq> 0\<close> by simp
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1025
  ultimately show ?Q by simp
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1026
qed
62376
85f38d5f8807 Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents: 62366
diff changeset
  1027
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1028
lemma div_normalize [simp]: "a div normalize a = unit_factor a"
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1029
proof (cases "a = 0")
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1030
  case True
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1031
  then show ?thesis by simp
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1032
next
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1033
  case False
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1034
  then have "normalize a \<noteq> 0" by simp
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1035
  with nonzero_mult_divide_cancel_right
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1036
  have "unit_factor a * normalize a div normalize a = unit_factor a" by blast
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1037
  then show ?thesis by simp
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1038
qed
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1039
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1040
lemma div_unit_factor [simp]: "a div unit_factor a = normalize a"
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1041
proof (cases "a = 0")
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1042
  case True
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1043
  then show ?thesis by simp
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1044
next
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1045
  case False
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1046
  then have "unit_factor a \<noteq> 0" by simp
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1047
  with nonzero_mult_divide_cancel_left
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1048
  have "unit_factor a * normalize a div unit_factor a = normalize a" by blast
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1049
  then show ?thesis by simp
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1050
qed
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1051
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1052
lemma normalize_div [simp]: "normalize a div a = 1 div unit_factor a"
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1053
proof (cases "a = 0")
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1054
  case True
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1055
  then show ?thesis by simp
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1056
next
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1057
  case False
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1058
  have "normalize a div a = normalize a div (unit_factor a * normalize a)"
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1059
    by simp
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1060
  also have "\<dots> = 1 div unit_factor a"
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1061
    using False by (subst is_unit_divide_mult_cancel_right) simp_all
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1062
  finally show ?thesis .
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1063
qed
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1064
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1065
lemma mult_one_div_unit_factor [simp]: "a * (1 div unit_factor b) = a div unit_factor b"
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1066
  by (cases "b = 0") simp_all
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1067
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1068
lemma normalize_mult: "normalize (a * b) = normalize a * normalize b"
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1069
proof (cases "a = 0 \<or> b = 0")
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1070
  case True
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1071
  then show ?thesis by auto
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1072
next
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1073
  case False
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1074
  from unit_factor_mult_normalize have "unit_factor (a * b) * normalize (a * b) = a * b" .
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1075
  then have "normalize (a * b) = a * b div unit_factor (a * b)"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1076
    by simp
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1077
  also have "\<dots> = a * b div unit_factor (b * a)"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1078
    by (simp add: ac_simps)
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1079
  also have "\<dots> = a * b div unit_factor b div unit_factor a"
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1080
    using False by (simp add: unit_factor_mult is_unit_div_mult2_eq [symmetric])
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1081
  also have "\<dots> = a * (b div unit_factor b) div unit_factor a"
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1082
    using False by (subst unit_div_mult_swap) simp_all
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1083
  also have "\<dots> = normalize a * normalize b"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1084
    using False
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1085
    by (simp add: mult.commute [of a] mult.commute [of "normalize a"] unit_div_mult_swap [symmetric])
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1086
  finally show ?thesis .
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1087
qed
62376
85f38d5f8807 Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents: 62366
diff changeset
  1088
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1089
lemma unit_factor_idem [simp]: "unit_factor (unit_factor a) = unit_factor a"
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1090
  by (cases "a = 0") (auto intro: is_unit_unit_factor)
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1091
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1092
lemma normalize_unit_factor [simp]: "a \<noteq> 0 \<Longrightarrow> normalize (unit_factor a) = 1"
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1093
  by (rule is_unit_normalize) simp
62376
85f38d5f8807 Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents: 62366
diff changeset
  1094
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1095
lemma normalize_idem [simp]: "normalize (normalize a) = normalize a"
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1096
proof (cases "a = 0")
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1097
  case True
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1098
  then show ?thesis by simp
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1099
next
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1100
  case False
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1101
  have "normalize a = normalize (unit_factor a * normalize a)"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1102
    by simp
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1103
  also have "\<dots> = normalize (unit_factor a) * normalize (normalize a)"
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1104
    by (simp only: normalize_mult)
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1105
  finally show ?thesis
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1106
    using False by simp_all
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1107
qed
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1108
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1109
lemma unit_factor_normalize [simp]:
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1110
  assumes "a \<noteq> 0"
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1111
  shows "unit_factor (normalize a) = 1"
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1112
proof -
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1113
  from assms have *: "normalize a \<noteq> 0"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1114
    by simp
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1115
  have "unit_factor (normalize a) * normalize (normalize a) = normalize a"
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1116
    by (simp only: unit_factor_mult_normalize)
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1117
  then have "unit_factor (normalize a) * normalize a = normalize a"
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1118
    by simp
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1119
  with * have "unit_factor (normalize a) * normalize a div normalize a = normalize a div normalize a"
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1120
    by simp
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1121
  with * show ?thesis
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1122
    by simp
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1123
qed
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1124
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1125
lemma dvd_unit_factor_div:
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1126
  assumes "b dvd a"
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1127
  shows "unit_factor (a div b) = unit_factor a div unit_factor b"
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1128
proof -
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1129
  from assms have "a = a div b * b"
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1130
    by simp
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1131
  then have "unit_factor a = unit_factor (a div b * b)"
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1132
    by simp
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1133
  then show ?thesis
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1134
    by (cases "b = 0") (simp_all add: unit_factor_mult)
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1135
qed
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1136
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1137
lemma dvd_normalize_div:
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1138
  assumes "b dvd a"
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1139
  shows "normalize (a div b) = normalize a div normalize b"
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1140
proof -
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1141
  from assms have "a = a div b * b"
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1142
    by simp
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1143
  then have "normalize a = normalize (a div b * b)"
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1144
    by simp
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1145
  then show ?thesis
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1146
    by (cases "b = 0") (simp_all add: normalize_mult)
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1147
qed
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1148
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1149
lemma normalize_dvd_iff [simp]: "normalize a dvd b \<longleftrightarrow> a dvd b"
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1150
proof -
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1151
  have "normalize a dvd b \<longleftrightarrow> unit_factor a * normalize a dvd b"
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1152
    using mult_unit_dvd_iff [of "unit_factor a" "normalize a" b]
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1153
      by (cases "a = 0") simp_all
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1154
  then show ?thesis by simp
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1155
qed
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1156
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1157
lemma dvd_normalize_iff [simp]: "a dvd normalize b \<longleftrightarrow> a dvd b"
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1158
proof -
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1159
  have "a dvd normalize  b \<longleftrightarrow> a dvd normalize b * unit_factor b"
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1160
    using dvd_mult_unit_iff [of "unit_factor b" a "normalize b"]
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1161
      by (cases "b = 0") simp_all
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1162
  then show ?thesis by simp
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1163
qed
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1164
60688
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1165
text \<open>
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1166
  We avoid an explicit definition of associated elements but prefer
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1167
  explicit normalisation instead.  In theory we could define an abbreviation
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1168
  like @{prop "associated a b \<longleftrightarrow> normalize a = normalize b"} but this is
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1169
  counterproductive without suggestive infix syntax, which we do not want
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1170
  to sacrifice for this purpose here.
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1171
\<close>
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1172
60688
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1173
lemma associatedI:
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1174
  assumes "a dvd b" and "b dvd a"
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1175
  shows "normalize a = normalize b"
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1176
proof (cases "a = 0 \<or> b = 0")
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1177
  case True
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1178
  with assms show ?thesis by auto
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1179
next
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1180
  case False
60688
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1181
  from \<open>a dvd b\<close> obtain c where b: "b = a * c" ..
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1182
  moreover from \<open>b dvd a\<close> obtain d where a: "a = b * d" ..
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1183
  ultimately have "b * 1 = b * (c * d)"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1184
    by (simp add: ac_simps)
60688
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1185
  with False have "1 = c * d"
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1186
    unfolding mult_cancel_left by simp
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1187
  then have "is_unit c" and "is_unit d"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1188
    by auto
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1189
  with a b show ?thesis
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1190
    by (simp add: normalize_mult is_unit_normalize)
60688
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1191
qed
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1192
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1193
lemma associatedD1: "normalize a = normalize b \<Longrightarrow> a dvd b"
60688
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1194
  using dvd_normalize_iff [of _ b, symmetric] normalize_dvd_iff [of a _, symmetric]
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1195
  by simp
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1196
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1197
lemma associatedD2: "normalize a = normalize b \<Longrightarrow> b dvd a"
60688
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1198
  using dvd_normalize_iff [of _ a, symmetric] normalize_dvd_iff [of b _, symmetric]
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1199
  by simp
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1200
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1201
lemma associated_unit: "normalize a = normalize b \<Longrightarrow> is_unit a \<Longrightarrow> is_unit b"
60688
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1202
  using dvd_unit_imp_unit by (auto dest!: associatedD1 associatedD2)
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1203
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1204
lemma associated_iff_dvd: "normalize a = normalize b \<longleftrightarrow> a dvd b \<and> b dvd a"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1205
  (is "?P \<longleftrightarrow> ?Q")
60688
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1206
proof
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1207
  assume ?Q
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1208
  then show ?P by (auto intro!: associatedI)
60688
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1209
next
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1210
  assume ?P
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1211
  then have "unit_factor a * normalize a = unit_factor a * normalize b"
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1212
    by simp
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1213
  then have *: "normalize b * unit_factor a = a"
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1214
    by (simp add: ac_simps)
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1215
  show ?Q
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1216
  proof (cases "a = 0 \<or> b = 0")
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1217
    case True
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1218
    with \<open>?P\<close> show ?thesis by auto
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1219
  next
62376
85f38d5f8807 Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents: 62366
diff changeset
  1220
    case False
60688
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1221
    then have "b dvd normalize b * unit_factor a" and "normalize b * unit_factor a dvd b"
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1222
      by (simp_all add: mult_unit_dvd_iff dvd_mult_unit_iff)
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1223
    with * show ?thesis by simp
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1224
  qed
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1225
qed
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1226
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1227
lemma associated_eqI:
60688
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1228
  assumes "a dvd b" and "b dvd a"
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1229
  assumes "normalize a = a" and "normalize b = b"
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1230
  shows "a = b"
60688
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1231
proof -
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1232
  from assms have "normalize a = normalize b"
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1233
    unfolding associated_iff_dvd by simp
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1234
  with \<open>normalize a = a\<close> have "a = normalize b" by simp
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1235
  with \<open>normalize b = b\<close> show "a = b" by simp
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1236
qed
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1237
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1238
end
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1239
62376
85f38d5f8807 Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents: 62366
diff changeset
  1240
class ordered_semiring = semiring + ordered_comm_monoid_add +
38642
8fa437809c67 dropped type classes mult_mono and mult_mono1; tuned names of technical rule duplicates
haftmann
parents: 37767
diff changeset
  1241
  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
  1242
  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
  1243
begin
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1244
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1245
lemma mult_mono: "a \<le> b \<Longrightarrow> c \<le> d \<Longrightarrow> 0 \<le> b \<Longrightarrow> 0 \<le> c \<Longrightarrow> a * c \<le> b * d"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1246
  apply (erule (1) mult_right_mono [THEN order_trans])
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1247
  apply (erule (1) mult_left_mono)
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1248
  done
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1249
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1250
lemma mult_mono': "a \<le> b \<Longrightarrow> c \<le> d \<Longrightarrow> 0 \<le> a \<Longrightarrow> 0 \<le> c \<Longrightarrow> a * c \<le> b * d"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1251
  apply (rule mult_mono)
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1252
  apply (fast intro: order_trans)+
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1253
  done
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1254
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1255
end
21199
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
  1256
62377
ace69956d018 moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
hoelzl
parents: 62376
diff changeset
  1257
class ordered_semiring_0 = semiring_0 + ordered_semiring
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
  1258
begin
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1259
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1260
lemma mult_nonneg_nonneg [simp]: "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> 0 \<le> a * b"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1261
  using mult_left_mono [of 0 b a] by simp
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1262
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1263
lemma mult_nonneg_nonpos: "0 \<le> a \<Longrightarrow> b \<le> 0 \<Longrightarrow> a * b \<le> 0"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1264
  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
  1265
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
  1266
lemma mult_nonpos_nonneg: "a \<le> 0 \<Longrightarrow> 0 \<le> b \<Longrightarrow> a * b \<le> 0"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1267
  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
  1268
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61762
diff changeset
  1269
text \<open>Legacy - use \<open>mult_nonpos_nonneg\<close>\<close>
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
  1270
lemma mult_nonneg_nonpos2: "0 \<le> a \<Longrightarrow> b \<le> 0 \<Longrightarrow> b * a \<le> 0"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1271
  apply (drule mult_right_mono [of b 0])
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1272
  apply auto
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1273
  done
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1274
62378
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  1275
lemma split_mult_neg_le: "(0 \<le> a \<and> b \<le> 0) \<or> (a \<le> 0 \<and> 0 \<le> b) \<Longrightarrow> a * b \<le> 0"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1276
  by (auto simp add: mult_nonneg_nonpos mult_nonneg_nonpos2)
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1277
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1278
end
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1279
62377
ace69956d018 moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
hoelzl
parents: 62376
diff changeset
  1280
class ordered_cancel_semiring = ordered_semiring + cancel_comm_monoid_add
ace69956d018 moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
hoelzl
parents: 62376
diff changeset
  1281
begin
ace69956d018 moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
hoelzl
parents: 62376
diff changeset
  1282
ace69956d018 moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
hoelzl
parents: 62376
diff changeset
  1283
subclass semiring_0_cancel ..
ace69956d018 moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
hoelzl
parents: 62376
diff changeset
  1284
subclass ordered_semiring_0 ..
ace69956d018 moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
hoelzl
parents: 62376
diff changeset
  1285
ace69956d018 moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
hoelzl
parents: 62376
diff changeset
  1286
end
ace69956d018 moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
hoelzl
parents: 62376
diff changeset
  1287
38642
8fa437809c67 dropped type classes mult_mono and mult_mono1; tuned names of technical rule duplicates
haftmann
parents: 37767
diff changeset
  1288
class linordered_semiring = ordered_semiring + linordered_cancel_ab_semigroup_add
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
  1289
begin
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1290
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1291
subclass ordered_cancel_semiring ..
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1292
62376
85f38d5f8807 Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents: 62366
diff changeset
  1293
subclass ordered_cancel_comm_monoid_add ..
25304
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  1294
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1295
lemma mult_left_less_imp_less: "c * a < c * b \<Longrightarrow> 0 \<le> c \<Longrightarrow> a < b"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1296
  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
  1297
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1298
lemma mult_right_less_imp_less: "a * c < b * c \<Longrightarrow> 0 \<le> c \<Longrightarrow> a < b"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1299
  by (force simp add: mult_right_mono not_le [symmetric])
23521
195fe3fe2831 added ordered_ring and ordered_semiring
obua
parents: 23496
diff changeset
  1300
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1301
lemma less_eq_add_cancel_left_greater_eq_zero [simp]: "a \<le> a + b \<longleftrightarrow> 0 \<le> b"
62608
19f87fa0cfcb more theorems on orderings
haftmann
parents: 62481
diff changeset
  1302
  using add_le_cancel_left [of a 0 b] by simp
19f87fa0cfcb more theorems on orderings
haftmann
parents: 62481
diff changeset
  1303
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1304
lemma less_eq_add_cancel_left_less_eq_zero [simp]: "a + b \<le> a \<longleftrightarrow> b \<le> 0"
62608
19f87fa0cfcb more theorems on orderings
haftmann
parents: 62481
diff changeset
  1305
  using add_le_cancel_left [of a b 0] by simp
19f87fa0cfcb more theorems on orderings
haftmann
parents: 62481
diff changeset
  1306
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1307
lemma less_eq_add_cancel_right_greater_eq_zero [simp]: "a \<le> b + a \<longleftrightarrow> 0 \<le> b"
62608
19f87fa0cfcb more theorems on orderings
haftmann
parents: 62481
diff changeset
  1308
  using add_le_cancel_right [of 0 a b] by simp
19f87fa0cfcb more theorems on orderings
haftmann
parents: 62481
diff changeset
  1309
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1310
lemma less_eq_add_cancel_right_less_eq_zero [simp]: "b + a \<le> a \<longleftrightarrow> b \<le> 0"
62608
19f87fa0cfcb more theorems on orderings
haftmann
parents: 62481
diff changeset
  1311
  using add_le_cancel_right [of b a 0] by simp
19f87fa0cfcb more theorems on orderings
haftmann
parents: 62481
diff changeset
  1312
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1313
lemma less_add_cancel_left_greater_zero [simp]: "a < a + b \<longleftrightarrow> 0 < b"
62608
19f87fa0cfcb more theorems on orderings
haftmann
parents: 62481
diff changeset
  1314
  using add_less_cancel_left [of a 0 b] by simp
19f87fa0cfcb more theorems on orderings
haftmann
parents: 62481
diff changeset
  1315
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1316
lemma less_add_cancel_left_less_zero [simp]: "a + b < a \<longleftrightarrow> b < 0"
62608
19f87fa0cfcb more theorems on orderings
haftmann
parents: 62481
diff changeset
  1317
  using add_less_cancel_left [of a b 0] by simp
19f87fa0cfcb more theorems on orderings
haftmann
parents: 62481
diff changeset
  1318
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1319
lemma less_add_cancel_right_greater_zero [simp]: "a < b + a \<longleftrightarrow> 0 < b"
62608
19f87fa0cfcb more theorems on orderings
haftmann
parents: 62481
diff changeset
  1320
  using add_less_cancel_right [of 0 a b] by simp
19f87fa0cfcb more theorems on orderings
haftmann
parents: 62481
diff changeset
  1321
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1322
lemma less_add_cancel_right_less_zero [simp]: "b + a < a \<longleftrightarrow> b < 0"
62608
19f87fa0cfcb more theorems on orderings
haftmann
parents: 62481
diff changeset
  1323
  using add_less_cancel_right [of b a 0] by simp
19f87fa0cfcb more theorems on orderings
haftmann
parents: 62481
diff changeset
  1324
25186
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
  1325
end
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
  1326
35043
07dbdf60d5ad dropped accidental duplication of "lin" prefix from cs. 108662d50512
haftmann
parents: 35032
diff changeset
  1327
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
  1328
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
  1329
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
  1330
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
  1331
  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
  1332
  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
  1333
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
  1334
  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
  1335
    by (simp add: add_mono mult_left_mono)
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1336
  with assms show ?thesis
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1337
    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
  1338
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
  1339
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
  1340
end
35043
07dbdf60d5ad dropped accidental duplication of "lin" prefix from cs. 108662d50512
haftmann
parents: 35032
diff changeset
  1341
07dbdf60d5ad dropped accidental duplication of "lin" prefix from cs. 108662d50512
haftmann
parents: 35032
diff changeset
  1342
class linordered_semiring_strict = semiring + comm_monoid_add + linordered_cancel_ab_semigroup_add +
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
  1343
  assumes mult_strict_left_mono: "a < b \<Longrightarrow> 0 < c \<Longrightarrow> c * a < c * b"
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
  1344
  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
  1345
begin
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14334
diff changeset
  1346
27516
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
  1347
subclass semiring_0_cancel ..
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
  1348
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1349
subclass linordered_semiring
28823
dcbef866c9e2 tuned unfold_locales invocation
haftmann
parents: 28559
diff changeset
  1350
proof
23550
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
  1351
  fix a b c :: 'a
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
  1352
  assume A: "a \<le> b" "0 \<le> c"
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
  1353
  from A show "c * a \<le> c * b"
25186
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
  1354
    unfolding le_less
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
  1355
    using mult_strict_left_mono by (cases "c = 0") auto
23550
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
  1356
  from A show "a * c \<le> b * c"
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
  1357
    unfolding le_less
25186
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
  1358
    using mult_strict_right_mono by (cases "c = 0") auto
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
  1359
qed
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
  1360
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1361
lemma mult_left_le_imp_le: "c * a \<le> c * b \<Longrightarrow> 0 < c \<Longrightarrow> a \<le> b"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1362
  by (auto 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
  1363
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1364
lemma mult_right_le_imp_le: "a * c \<le> b * c \<Longrightarrow> 0 < c \<Longrightarrow> a \<le> b"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1365
  by (auto simp add: mult_strict_right_mono not_less [symmetric])
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1366
56544
b60d5d119489 made mult_pos_pos a simp rule
nipkow
parents: 56536
diff changeset
  1367
lemma mult_pos_pos[simp]: "0 < a \<Longrightarrow> 0 < b \<Longrightarrow> 0 < a * b"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1368
  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
  1369
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
  1370
lemma mult_pos_neg: "0 < a \<Longrightarrow> b < 0 \<Longrightarrow> a * b < 0"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1371
  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
  1372
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
  1373
lemma mult_neg_pos: "a < 0 \<Longrightarrow> 0 < b \<Longrightarrow> a * b < 0"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1374
  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
  1375
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61762
diff changeset
  1376
text \<open>Legacy - use \<open>mult_neg_pos\<close>\<close>
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
  1377
lemma mult_pos_neg2: "0 < a \<Longrightarrow> b < 0 \<Longrightarrow> b * a < 0"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1378
  apply (drule mult_strict_right_mono [of b 0])
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1379
  apply auto
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1380
  done
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1381
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1382
lemma zero_less_mult_pos: "0 < a * b \<Longrightarrow> 0 < a \<Longrightarrow> 0 < b"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1383
  apply (cases "b \<le> 0")
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1384
   apply (auto simp add: le_less not_less)
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1385
  apply (drule_tac mult_pos_neg [of a b])
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1386
   apply (auto dest: less_not_sym)
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1387
  done
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1388
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1389
lemma zero_less_mult_pos2: "0 < b * a \<Longrightarrow> 0 < a \<Longrightarrow> 0 < b"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1390
  apply (cases "b \<le> 0")
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1391
   apply (auto simp add: le_less not_less)
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1392
  apply (drule_tac mult_pos_neg2 [of a b])
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1393
   apply (auto dest: less_not_sym)
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1394
  done
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1395
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1396
text \<open>Strict monotonicity in both arguments\<close>
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1397
lemma mult_strict_mono:
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1398
  assumes "a < b" and "c < d" and "0 < b" and "0 \<le> c"
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1399
  shows "a * c < b * d"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1400
  using assms
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1401
  apply (cases "c = 0")
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1402
  apply simp
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1403
  apply (erule mult_strict_right_mono [THEN less_trans])
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1404
  apply (auto simp add: le_less)
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1405
  apply (erule (1) mult_strict_left_mono)
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1406
  done
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1407
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1408
text \<open>This weaker variant has more natural premises\<close>
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1409
lemma mult_strict_mono':
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1410
  assumes "a < b" and "c < d" and "0 \<le> a" and "0 \<le> c"
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1411
  shows "a * c < b * d"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1412
  by (rule mult_strict_mono) (insert assms, auto)
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1413
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1414
lemma mult_less_le_imp_less:
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1415
  assumes "a < b" and "c \<le> d" and "0 \<le> a" and "0 < c"
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1416
  shows "a * c < b * d"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1417
  using assms
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1418
  apply (subgoal_tac "a * c < b * c")
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1419
  apply (erule less_le_trans)
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1420
  apply (erule mult_left_mono)
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1421
  apply simp
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1422
  apply (erule (1) mult_strict_right_mono)
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1423
  done
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1424
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1425
lemma mult_le_less_imp_less:
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1426
  assumes "a \<le> b" and "c < d" and "0 < a" and "0 \<le> c"
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1427
  shows "a * c < b * d"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1428
  using assms
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1429
  apply (subgoal_tac "a * c \<le> b * c")
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1430
  apply (erule le_less_trans)
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1431
  apply (erule mult_strict_left_mono)
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1432
  apply simp
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1433
  apply (erule (1) mult_right_mono)
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1434
  done
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1435
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1436
end
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1437
35097
4554bb2abfa3 dropped last occurence of the linlinordered accident
haftmann
parents: 35092
diff changeset
  1438
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
  1439
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
  1440
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
  1441
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
  1442
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
  1443
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
  1444
  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
  1445
  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
  1446
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
  1447
  from assms have "u * x + v * y < u * a + v * a"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1448
    by (cases "u = 0") (auto intro!: add_less_le_mono mult_strict_left_mono mult_left_mono)
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1449
  with assms show ?thesis
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1450
    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
  1451
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
  1452
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
  1453
end
33319
74f0dcc0b5fb moved algebraic classes to Ring_and_Field
haftmann
parents: 32960
diff changeset
  1454
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
  1455
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
  1456
  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
  1457
begin
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
  1458
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1459
subclass ordered_semiring
28823
dcbef866c9e2 tuned unfold_locales invocation
haftmann
parents: 28559
diff changeset
  1460
proof
21199
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
  1461
  fix a b c :: 'a
23550
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
  1462
  assume "a \<le> b" "0 \<le> c"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1463
  then show "c * a \<le> c * b" by (rule comm_mult_left_mono)
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1464
  then show "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
  1465
qed
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
  1466
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
  1467
end
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
  1468
38642
8fa437809c67 dropped type classes mult_mono and mult_mono1; tuned names of technical rule duplicates
haftmann
parents: 37767
diff changeset
  1469
class ordered_cancel_comm_semiring = ordered_comm_semiring + cancel_comm_monoid_add
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
  1470
begin
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
  1471
38642
8fa437809c67 dropped type classes mult_mono and mult_mono1; tuned names of technical rule duplicates
haftmann
parents: 37767
diff changeset
  1472
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
  1473
subclass ordered_comm_semiring ..
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1474
subclass ordered_cancel_semiring ..
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
  1475
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
  1476
end
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
  1477
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1478
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
  1479
  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
  1480
begin
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
  1481
35043
07dbdf60d5ad dropped accidental duplication of "lin" prefix from cs. 108662d50512
haftmann
parents: 35032
diff changeset
  1482
subclass linordered_semiring_strict
28823
dcbef866c9e2 tuned unfold_locales invocation
haftmann
parents: 28559
diff changeset
  1483
proof
23550
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
  1484
  fix a b c :: 'a
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
  1485
  assume "a < b" "0 < c"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1486
  then show "c * a < c * b" by (rule comm_mult_strict_left_mono)
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1487
  then show "a * c < b * c" by (simp only: mult.commute)
23550
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
  1488
qed
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
  1489
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1490
subclass ordered_cancel_comm_semiring
28823
dcbef866c9e2 tuned unfold_locales invocation
haftmann
parents: 28559
diff changeset
  1491
proof
23550
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
  1492
  fix a b c :: 'a
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
  1493
  assume "a \<le> b" "0 \<le> c"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1494
  then show "c * a \<le> c * b"
25186
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
  1495
    unfolding le_less
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1496
    using mult_strict_left_mono by (cases "c = 0") auto
23550
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
  1497
qed
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
  1498
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
  1499
end
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1500
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
  1501
class ordered_ring = ring + ordered_cancel_semiring
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
  1502
begin
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1503
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1504
subclass ordered_ab_group_add ..
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1505
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1506
lemma less_add_iff1: "a * e + c < b * e + d \<longleftrightarrow> (a - b) * e + c < d"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1507
  by (simp add: algebra_simps)
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1508
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1509
lemma less_add_iff2: "a * e + c < b * e + d \<longleftrightarrow> c < (b - a) * e + d"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1510
  by (simp add: algebra_simps)
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1511
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1512
lemma le_add_iff1: "a * e + c \<le> b * e + d \<longleftrightarrow> (a - b) * e + c \<le> d"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1513
  by (simp add: algebra_simps)
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1514
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1515
lemma le_add_iff2: "a * e + c \<le> b * e + d \<longleftrightarrow> c \<le> (b - a) * e + d"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1516
  by (simp add: algebra_simps)
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1517
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1518
lemma mult_left_mono_neg: "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
  1519
  apply (drule mult_left_mono [of _ _ "- c"])
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35097
diff changeset
  1520
  apply simp_all
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1521
  done
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1522
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1523
lemma mult_right_mono_neg: "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
  1524
  apply (drule mult_right_mono [of _ _ "- c"])
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35097
diff changeset
  1525
  apply simp_all
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1526
  done
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1527
30692
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
  1528
lemma mult_nonpos_nonpos: "a \<le> 0 \<Longrightarrow> b \<le> 0 \<Longrightarrow> 0 \<le> a * b"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1529
  using mult_right_mono_neg [of a 0 b] by simp
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1530
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1531
lemma split_mult_pos_le: "(0 \<le> a \<and> 0 \<le> b) \<or> (a \<le> 0 \<and> b \<le> 0) \<Longrightarrow> 0 \<le> a * b"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1532
  by (auto simp add: mult_nonpos_nonpos)
25186
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
  1533
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
  1534
end
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1535
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1536
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
  1537
begin
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  1538
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1539
subclass ordered_ring ..
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1540
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1541
subclass ordered_ab_group_add_abs
28823
dcbef866c9e2 tuned unfold_locales invocation
haftmann
parents: 28559
diff changeset
  1542
proof
25304
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  1543
  fix a b
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  1544
  show "\<bar>a + b\<bar> \<le> \<bar>a\<bar> + \<bar>b\<bar>"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1545
    by (auto simp add: abs_if not_le not_less algebra_simps
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1546
        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
  1547
qed (auto simp add: abs_if)
25304
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  1548
35631
0b8a5fd339ab generalize some lemmas from class linordered_ring_strict to linordered_ring
huffman
parents: 35302
diff changeset
  1549
lemma zero_le_square [simp]: "0 \<le> a * a"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1550
  using linear [of 0 a] 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
  1551
0b8a5fd339ab generalize some lemmas from class linordered_ring_strict to linordered_ring
huffman
parents: 35302
diff changeset
  1552
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
  1553
  by (simp add: not_less)
0b8a5fd339ab generalize some lemmas from class linordered_ring_strict to linordered_ring
huffman
parents: 35302
diff changeset
  1554
61944
5d06ecfdb472 prefer symbols for "abs";
wenzelm
parents: 61799
diff changeset
  1555
proposition abs_eq_iff: "\<bar>x\<bar> = \<bar>y\<bar> \<longleftrightarrow> x = y \<or> x = -y"
62390
842917225d56 more canonical names
nipkow
parents: 62378
diff changeset
  1556
  by (auto simp add: abs_if split: if_split_asm)
61762
d50b993b4fb9 Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents: 61649
diff changeset
  1557
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1558
lemma sum_squares_ge_zero: "0 \<le> x * x + y * y"
62347
2230b7047376 generalized some lemmas;
haftmann
parents: 61944
diff changeset
  1559
  by (intro add_nonneg_nonneg zero_le_square)
2230b7047376 generalized some lemmas;
haftmann
parents: 61944
diff changeset
  1560
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1561
lemma not_sum_squares_lt_zero: "\<not> x * x + y * y < 0"
62347
2230b7047376 generalized some lemmas;
haftmann
parents: 61944
diff changeset
  1562
  by (simp add: not_less sum_squares_ge_zero)
2230b7047376 generalized some lemmas;
haftmann
parents: 61944
diff changeset
  1563
25304
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  1564
end
23521
195fe3fe2831 added ordered_ring and ordered_semiring
obua
parents: 23496
diff changeset
  1565
35043
07dbdf60d5ad dropped accidental duplication of "lin" prefix from cs. 108662d50512
haftmann
parents: 35032
diff changeset
  1566
class linordered_ring_strict = ring + linordered_semiring_strict
25304
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  1567
  + ordered_ab_group_add + abs_if
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1568
begin
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
  1569
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1570
subclass linordered_ring ..
25304
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  1571
30692
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
  1572
lemma mult_strict_left_mono_neg: "b < a \<Longrightarrow> c < 0 \<Longrightarrow> c * a < c * b"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1573
  using mult_strict_left_mono [of b a "- c"] by simp
30692
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
  1574
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
  1575
lemma mult_strict_right_mono_neg: "b < a \<Longrightarrow> c < 0 \<Longrightarrow> a * c < b * c"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1576
  using mult_strict_right_mono [of b a "- c"] by simp
30692
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
  1577
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
  1578
lemma mult_neg_neg: "a < 0 \<Longrightarrow> b < 0 \<Longrightarrow> 0 < a * b"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1579
  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
  1580
25917
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1581
subclass ring_no_zero_divisors
28823
dcbef866c9e2 tuned unfold_locales invocation
haftmann
parents: 28559
diff changeset
  1582
proof
25917
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1583
  fix a b
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1584
  assume "a \<noteq> 0"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1585
  then have A: "a < 0 \<or> 0 < a" by (simp add: neq_iff)
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1586
  assume "b \<noteq> 0"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1587
  then have B: "b < 0 \<or> 0 < b" by (simp add: neq_iff)
25917
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1588
  have "a * b < 0 \<or> 0 < a * b"
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1589
  proof (cases "a < 0")
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1590
    case A': True
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1591
    show ?thesis
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1592
    proof (cases "b < 0")
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1593
      case True
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1594
      with A' show ?thesis by (auto dest: mult_neg_neg)
25917
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1595
    next
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1596
      case False
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1597
      with B have "0 < b" by auto
25917
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1598
      with A' show ?thesis by (auto dest: mult_strict_right_mono)
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1599
    qed
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1600
  next
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1601
    case False
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1602
    with A have A': "0 < a" by auto
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1603
    show ?thesis
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1604
    proof (cases "b < 0")
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1605
      case True
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1606
      with A' show ?thesis
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1607
        by (auto dest: mult_strict_right_mono_neg)
25917
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1608
    next
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1609
      case False
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1610
      with B have "0 < b" by auto
56544
b60d5d119489 made mult_pos_pos a simp rule
nipkow
parents: 56536
diff changeset
  1611
      with A' show ?thesis by auto
25917
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1612
    qed
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1613
  qed
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1614
  then show "a * b \<noteq> 0"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1615
    by (simp add: neq_iff)
25917
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1616
qed
25304
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  1617
56480
093ea91498e6 field_simps: better support for negation and division, and power
hoelzl
parents: 56217
diff changeset
  1618
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
  1619
  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
  1620
     (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
  1621
56480
093ea91498e6 field_simps: better support for negation and division, and power
hoelzl
parents: 56217
diff changeset
  1622
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
  1623
  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
  1624
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1625
lemma mult_less_0_iff: "a * b < 0 \<longleftrightarrow> 0 < a \<and> b < 0 \<or> a < 0 \<and> 0 < b"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1626
  using zero_less_mult_iff [of "- a" b] by auto
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
  1627
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1628
lemma mult_le_0_iff: "a * b \<le> 0 \<longleftrightarrow> 0 \<le> a \<and> b \<le> 0 \<or> a \<le> 0 \<and> 0 \<le> b"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1629
  using zero_le_mult_iff [of "- a" b] by auto
25917
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1630
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1631
text \<open>
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1632
  Cancellation laws for @{term "c * a < c * b"} and @{term "a * c < b * c"},
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1633
  also with the relations \<open>\<le>\<close> and equality.
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1634
\<close>
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1635
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1636
text \<open>
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1637
  These ``disjunction'' versions produce two cases when the comparison is
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1638
  an assumption, but effectively four when the comparison is a goal.
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1639
\<close>
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1640
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1641
lemma mult_less_cancel_right_disj: "a * c < b * c \<longleftrightarrow> 0 < c \<and> a < b \<or> c < 0 \<and>  b < a"
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1642
  apply (cases "c = 0")
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1643
  apply (auto simp add: neq_iff mult_strict_right_mono mult_strict_right_mono_neg)
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1644
  apply (auto simp add: not_less not_le [symmetric, of "a*c"] not_le [symmetric, of a])
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1645
  apply (erule_tac [!] notE)
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1646
  apply (auto simp add: less_imp_le mult_right_mono mult_right_mono_neg)
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1647
  done
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1648
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1649
lemma mult_less_cancel_left_disj: "c * a < c * b \<longleftrightarrow> 0 < c \<and> a < b \<or> c < 0 \<and>  b < a"
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1650
  apply (cases "c = 0")
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1651
  apply (auto simp add: neq_iff mult_strict_left_mono mult_strict_left_mono_neg)
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1652
  apply (auto simp add: not_less not_le [symmetric, of "c * a"] not_le [symmetric, of a])
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1653
  apply (erule_tac [!] notE)
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1654
  apply (auto simp add: less_imp_le mult_left_mono mult_left_mono_neg)
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1655
  done
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1656
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1657
text \<open>
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1658
  The ``conjunction of implication'' lemmas produce two cases when the
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1659
  comparison is a goal, but give four when the comparison is an assumption.
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1660
\<close>
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1661
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1662
lemma mult_less_cancel_right: "a * c < b * c \<longleftrightarrow> (0 \<le> c \<longrightarrow> a < b) \<and> (c \<le> 0 \<longrightarrow> b < a)"
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1663
  using mult_less_cancel_right_disj [of a c b] by auto
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1664
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1665
lemma mult_less_cancel_left: "c * a < c * b \<longleftrightarrow> (0 \<le> c \<longrightarrow> a < b) \<and> (c \<le> 0 \<longrightarrow> b < a)"
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1666
  using mult_less_cancel_left_disj [of c a b] by auto
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1667
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1668
lemma mult_le_cancel_right: "a * c \<le> b * c \<longleftrightarrow> (0 < c \<longrightarrow> a \<le> b) \<and> (c < 0 \<longrightarrow> b \<le> a)"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1669
  by (simp add: not_less [symmetric] mult_less_cancel_right_disj)
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1670
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1671
lemma mult_le_cancel_left: "c * a \<le> c * b \<longleftrightarrow> (0 < c \<longrightarrow> a \<le> b) \<and> (c < 0 \<longrightarrow> b \<le> a)"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1672
  by (simp add: not_less [symmetric] mult_less_cancel_left_disj)
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1673
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1674
lemma mult_le_cancel_left_pos: "0 < c \<Longrightarrow> c * a \<le> c * b \<longleftrightarrow> a \<le> b"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1675
  by (auto simp: mult_le_cancel_left)
30649
57753e0ec1d4 1. New cancellation simprocs for common factors in inequations
nipkow
parents: 30242
diff changeset
  1676
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1677
lemma mult_le_cancel_left_neg: "c < 0 \<Longrightarrow> c * a \<le> c * b \<longleftrightarrow> b \<le> a"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1678
  by (auto simp: mult_le_cancel_left)
30649
57753e0ec1d4 1. New cancellation simprocs for common factors in inequations
nipkow
parents: 30242
diff changeset
  1679
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1680
lemma mult_less_cancel_left_pos: "0 < c \<Longrightarrow> c * a < c * b \<longleftrightarrow> a < b"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1681
  by (auto simp: mult_less_cancel_left)
30649
57753e0ec1d4 1. New cancellation simprocs for common factors in inequations
nipkow
parents: 30242
diff changeset
  1682
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1683
lemma mult_less_cancel_left_neg: "c < 0 \<Longrightarrow> c * a < c * b \<longleftrightarrow> b < a"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1684
  by (auto simp: mult_less_cancel_left)
30649
57753e0ec1d4 1. New cancellation simprocs for common factors in inequations
nipkow
parents: 30242
diff changeset
  1685
25917
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1686
end
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
  1687
30692
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
  1688
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
  1689
  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
  1690
  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
  1691
  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
  1692
  mult_neg_pos mult_neg_neg
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1693
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1694
class ordered_comm_ring = comm_ring + ordered_comm_semiring
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
  1695
begin
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1696
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1697
subclass ordered_ring ..
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1698
subclass ordered_cancel_comm_semiring ..
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1699
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
  1700
end
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1701
62378
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  1702
class zero_less_one = order + zero + one +
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1703
  assumes zero_less_one [simp]: "0 < 1"
62378
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  1704
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  1705
class linordered_nonzero_semiring = ordered_comm_semiring + monoid_mult + linorder + zero_less_one
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  1706
begin
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  1707
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  1708
subclass zero_neq_one
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1709
  by standard (insert zero_less_one, blast)
62378
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  1710
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  1711
subclass comm_semiring_1
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1712
  by standard (rule mult_1_left)
62378
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  1713
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  1714
lemma zero_le_one [simp]: "0 \<le> 1"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1715
  by (rule zero_less_one [THEN less_imp_le])
62378
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  1716
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  1717
lemma not_one_le_zero [simp]: "\<not> 1 \<le> 0"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1718
  by (simp add: not_le)
62378
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  1719
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  1720
lemma not_one_less_zero [simp]: "\<not> 1 < 0"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1721
  by (simp add: not_less)
62378
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  1722
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  1723
lemma mult_left_le: "c \<le> 1 \<Longrightarrow> 0 \<le> a \<Longrightarrow> a * c \<le> a"
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  1724
  using mult_left_mono[of c 1 a] by simp
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  1725
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  1726
lemma mult_le_one: "a \<le> 1 \<Longrightarrow> 0 \<le> b \<Longrightarrow> b \<le> 1 \<Longrightarrow> a * b \<le> 1"
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  1727
  using mult_mono[of a 1 b 1] by simp
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  1728
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  1729
lemma zero_less_two: "0 < 1 + 1"
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  1730
  using add_pos_pos[OF zero_less_one zero_less_one] .
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  1731
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  1732
end
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  1733
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  1734
class linordered_semidom = semidom + linordered_comm_semiring_strict + zero_less_one +
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
  1735
  assumes le_add_diff_inverse2 [simp]: "b \<le> a \<Longrightarrow> a - b + b = a"
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1736
begin
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1737
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1738
subclass linordered_nonzero_semiring ..
62378
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  1739
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
  1740
text \<open>Addition is the inverse of subtraction.\<close>
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
  1741
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
  1742
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
  1743
  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
  1744
62378
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  1745
lemma add_diff_inverse: "\<not> a < b \<Longrightarrow> b + (a - b) = 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
  1746
  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
  1747
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1748
lemma add_le_imp_le_diff: "i + k \<le> n \<Longrightarrow> i \<le> n - k"
60615
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 60570
diff changeset
  1749
  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
  1750
  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
  1751
  apply (simp only: add.assoc [symmetric])
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1752
  using add_implies_diff apply fastforce
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1753
  done
60615
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 60570
diff changeset
  1754
62376
85f38d5f8807 Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents: 62366
diff changeset
  1755
lemma add_le_add_imp_diff_le:
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1756
  assumes 1: "i + k \<le> n"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1757
    and 2: "n \<le> j + k"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1758
  shows "i + k \<le> n \<Longrightarrow> n \<le> j + k \<Longrightarrow> n - k \<le> j"
60615
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 60570
diff changeset
  1759
proof -
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 60570
diff changeset
  1760
  have "n - (i + k) + (i + k) = n"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1761
    using 1 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
  1762
  moreover have "n - k = n - k - i + i"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1763
    using 1 by (simp add: add_le_imp_le_diff)
60615
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 60570
diff changeset
  1764
  ultimately show ?thesis
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1765
    using 2
60615
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 60570
diff changeset
  1766
    apply (simp add: add.assoc [symmetric])
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1767
    apply (rule add_le_imp_le_diff [of _ k "j + k", simplified add_diff_cancel_right'])
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1768
    apply (simp add: add.commute diff_diff_add)
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1769
    done
60615
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 60570
diff changeset
  1770
qed
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 60570
diff changeset
  1771
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1772
lemma less_1_mult: "1 < m \<Longrightarrow> 1 < n \<Longrightarrow> 1 < m * n"
62378
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  1773
  using mult_strict_mono [of 1 m 1 n] by (simp add: less_trans [OF zero_less_one])
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58952
diff changeset
  1774
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1775
end
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1776
62378
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  1777
class linordered_idom =
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  1778
  comm_ring_1 + linordered_comm_semiring_strict + ordered_ab_group_add + abs_if + sgn_if
25917
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1779
begin
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1780
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
  1781
subclass linordered_semiring_1_strict ..
35043
07dbdf60d5ad dropped accidental duplication of "lin" prefix from cs. 108662d50512
haftmann
parents: 35032
diff changeset
  1782
subclass linordered_ring_strict ..
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1783
subclass ordered_comm_ring ..
27516
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
  1784
subclass idom ..
25917
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1785
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1786
subclass linordered_semidom
28823
dcbef866c9e2 tuned unfold_locales invocation
haftmann
parents: 28559
diff changeset
  1787
proof
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1788
  have "0 \<le> 1 * 1" by (rule zero_le_square)
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1789
  then show "0 < 1" by (simp add: le_less)
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1790
  show "b \<le> a \<Longrightarrow> a - b + b = a" for 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
  1791
    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
  1792
qed
25917
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1793
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1794
lemma linorder_neqE_linordered_idom:
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1795
  assumes "x \<noteq> y"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1796
  obtains "x < y" | "y < x"
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1797
  using assms by (rule neqE)
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1798
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
  1799
text \<open>These cancellation simprules also produce two cases when the comparison is a goal.\<close>
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  1800
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1801
lemma mult_le_cancel_right1: "c \<le> b * c \<longleftrightarrow> (0 < c \<longrightarrow> 1 \<le> b) \<and> (c < 0 \<longrightarrow> b \<le> 1)"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1802
  using mult_le_cancel_right [of 1 c b] by simp
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  1803
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1804
lemma mult_le_cancel_right2: "a * c \<le> c \<longleftrightarrow> (0 < c \<longrightarrow> a \<le> 1) \<and> (c < 0 \<longrightarrow> 1 \<le> a)"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1805
  using mult_le_cancel_right [of a c 1] by simp
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  1806
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1807
lemma mult_le_cancel_left1: "c \<le> c * b \<longleftrightarrow> (0 < c \<longrightarrow> 1 \<le> b) \<and> (c < 0 \<longrightarrow> b \<le> 1)"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1808
  using mult_le_cancel_left [of c 1 b] by simp
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  1809
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1810
lemma mult_le_cancel_left2: "c * a \<le> c \<longleftrightarrow> (0 < c \<longrightarrow> a \<le> 1) \<and> (c < 0 \<longrightarrow> 1 \<le> a)"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1811
  using mult_le_cancel_left [of c a 1] by simp
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  1812
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1813
lemma mult_less_cancel_right1: "c < b * c \<longleftrightarrow> (0 \<le> c \<longrightarrow> 1 < b) \<and> (c \<le> 0 \<longrightarrow> b < 1)"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1814
  using mult_less_cancel_right [of 1 c b] by simp
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  1815
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1816
lemma mult_less_cancel_right2: "a * c < c \<longleftrightarrow> (0 \<le> c \<longrightarrow> a < 1) \<and> (c \<le> 0 \<longrightarrow> 1 < a)"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1817
  using mult_less_cancel_right [of a c 1] by simp
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  1818
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1819
lemma mult_less_cancel_left1: "c < c * b \<longleftrightarrow> (0 \<le> c \<longrightarrow> 1 < b) \<and> (c \<le> 0 \<longrightarrow> b < 1)"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1820
  using mult_less_cancel_left [of c 1 b] by simp
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  1821
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1822
lemma mult_less_cancel_left2: "c * a < c \<longleftrightarrow> (0 \<le> c \<longrightarrow> a < 1) \<and> (c \<le> 0 \<longrightarrow> 1 < a)"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1823
  using mult_less_cancel_left [of c a 1] by simp
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  1824
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1825
lemma sgn_sgn [simp]: "sgn (sgn a) = sgn a"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1826
  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
  1827
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1828
lemma sgn_0_0: "sgn a = 0 \<longleftrightarrow> a = 0"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1829
  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
  1830
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1831
lemma sgn_1_pos: "sgn a = 1 \<longleftrightarrow> a > 0"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1832
  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
  1833
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1834
lemma sgn_1_neg: "sgn a = - 1 \<longleftrightarrow> a < 0"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1835
  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
  1836
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1837
lemma sgn_pos [simp]: "0 < a \<Longrightarrow> sgn a = 1"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1838
  by (simp only: sgn_1_pos)
29940
83b373f61d41 more default simp rules for sgn
haftmann
parents: 29925
diff changeset
  1839
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1840
lemma sgn_neg [simp]: "a < 0 \<Longrightarrow> sgn a = - 1"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1841
  by (simp only: sgn_1_neg)
29940
83b373f61d41 more default simp rules for sgn
haftmann
parents: 29925
diff changeset
  1842
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1843
lemma sgn_times: "sgn (a * b) = sgn a * sgn b"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1844
  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
  1845
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1846
lemma abs_sgn: "\<bar>k\<bar> = k * sgn k"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1847
  unfolding sgn_if abs_if by auto
29700
22faf21db3df added some simp rules
nipkow
parents: 29668
diff changeset
  1848
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1849
lemma sgn_greater [simp]: "0 < sgn a \<longleftrightarrow> 0 < a"
29940
83b373f61d41 more default simp rules for sgn
haftmann
parents: 29925
diff changeset
  1850
  unfolding sgn_if by auto
83b373f61d41 more default simp rules for sgn
haftmann
parents: 29925
diff changeset
  1851
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1852
lemma sgn_less [simp]: "sgn a < 0 \<longleftrightarrow> a < 0"
29940
83b373f61d41 more default simp rules for sgn
haftmann
parents: 29925
diff changeset
  1853
  unfolding sgn_if by auto
83b373f61d41 more default simp rules for sgn
haftmann
parents: 29925
diff changeset
  1854
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1855
lemma abs_sgn_eq: "\<bar>sgn a\<bar> = (if a = 0 then 0 else 1)"
62347
2230b7047376 generalized some lemmas;
haftmann
parents: 61944
diff changeset
  1856
  by (simp add: sgn_if)
2230b7047376 generalized some lemmas;
haftmann
parents: 61944
diff changeset
  1857
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1858
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
  1859
  by (simp add: abs_if)
20a506b8256d lemmas abs_dvd_iff, dvd_abs_iff
huffman
parents: 29940
diff changeset
  1860
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1861
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
  1862
  by (simp add: abs_if)
29653
ece6a0e9f8af added lemma abs_sng
haftmann
parents: 29465
diff changeset
  1863
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1864
lemma dvd_if_abs_eq: "\<bar>l\<bar> = \<bar>k\<bar> \<Longrightarrow> l dvd k"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1865
  by (subst abs_dvd_iff [symmetric]) simp
33676
802f5e233e48 moved lemma from Algebra/IntRing to Ring_and_Field
nipkow
parents: 33364
diff changeset
  1866
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1867
text \<open>
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1868
  The following lemmas can be proven in more general structures, but
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1869
  are dangerous as simp rules in absence of @{thm neg_equal_zero},
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1870
  @{thm neg_less_pos}, @{thm neg_less_eq_nonneg}.
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1871
\<close>
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  1872
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1873
lemma equation_minus_iff_1 [simp, no_atp]: "1 = - a \<longleftrightarrow> a = - 1"
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  1874
  by (fact equation_minus_iff)
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  1875
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1876
lemma minus_equation_iff_1 [simp, no_atp]: "- a = 1 \<longleftrightarrow> a = - 1"
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  1877
  by (subst minus_equation_iff, auto)
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  1878
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1879
lemma le_minus_iff_1 [simp, no_atp]: "1 \<le> - b \<longleftrightarrow> b \<le> - 1"
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  1880
  by (fact le_minus_iff)
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  1881
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1882
lemma minus_le_iff_1 [simp, no_atp]: "- a \<le> 1 \<longleftrightarrow> - 1 \<le> a"
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  1883
  by (fact minus_le_iff)
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  1884
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1885
lemma less_minus_iff_1 [simp, no_atp]: "1 < - b \<longleftrightarrow> b < - 1"
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  1886
  by (fact less_minus_iff)
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  1887
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1888
lemma minus_less_iff_1 [simp, no_atp]: "- a < 1 \<longleftrightarrow> - 1 < a"
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  1889
  by (fact minus_less_iff)
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  1890
25917
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1891
end
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1892
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
  1893
text \<open>Simprules for comparisons where common factors can be cancelled.\<close>
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
  1894
54147
97a8ff4e4ac9 killed most "no_atp", to make Sledgehammer more complete
blanchet
parents: 52435
diff changeset
  1895
lemmas mult_compare_simps =
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1896
  mult_le_cancel_right mult_le_cancel_left
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1897
  mult_le_cancel_right1 mult_le_cancel_right2
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1898
  mult_le_cancel_left1 mult_le_cancel_left2
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1899
  mult_less_cancel_right mult_less_cancel_left
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1900
  mult_less_cancel_right1 mult_less_cancel_right2
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1901
  mult_less_cancel_left1 mult_less_cancel_left2
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1902
  mult_cancel_right mult_cancel_left
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1903
  mult_cancel_right1 mult_cancel_right2
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1904
  mult_cancel_left1 mult_cancel_left2
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1905
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
  1906
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
  1907
text \<open>Reasoning about inequalities with division\<close>
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1908
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1909
context linordered_semidom
25193
e2e1a4b00de3 various localizations
haftmann
parents: 25186
diff changeset
  1910
begin
e2e1a4b00de3 various localizations
haftmann
parents: 25186
diff changeset
  1911
e2e1a4b00de3 various localizations
haftmann
parents: 25186
diff changeset
  1912
lemma less_add_one: "a < a + 1"
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1913
proof -
25193
e2e1a4b00de3 various localizations
haftmann
parents: 25186
diff changeset
  1914
  have "a + 0 < a + 1"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1915
    by (blast intro: zero_less_one add_strict_left_mono)
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1916
  then show ?thesis by simp
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1917
qed
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1918
25193
e2e1a4b00de3 various localizations
haftmann
parents: 25186
diff changeset
  1919
end
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1920
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1921
context linordered_idom
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1922
begin
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
  1923
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1924
lemma mult_right_le_one_le: "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
  1925
  by (rule mult_left_le)
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1926
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1927
lemma mult_left_le_one_le: "0 \<le> x \<Longrightarrow> 0 \<le> y \<Longrightarrow> y \<le> 1 \<Longrightarrow> y * x \<le> x"
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1928
  by (auto simp add: mult_le_cancel_right2)
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1929
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1930
end
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1931
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
  1932
text \<open>Absolute Value\<close>
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1933
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1934
context linordered_idom
25304
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  1935
begin
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  1936
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1937
lemma mult_sgn_abs: "sgn x * \<bar>x\<bar> = x"
25304
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  1938
  unfolding abs_if sgn_if by auto
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  1939
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1940
lemma abs_one [simp]: "\<bar>1\<bar> = 1"
44921
58eef4843641 tuned proofs
huffman
parents: 44350
diff changeset
  1941
  by (simp add: abs_if)
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1942
25304
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  1943
end
24491
8d194c9198ae added constant sgn
nipkow
parents: 24427
diff changeset
  1944
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1945
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
  1946
  assumes abs_eq_mult:
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  1947
    "(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
  1948
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1949
context linordered_idom
30961
541bfff659af more localisation
haftmann
parents: 30692
diff changeset
  1950
begin
541bfff659af more localisation
haftmann
parents: 30692
diff changeset
  1951
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1952
subclass ordered_ring_abs
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1953
  by standard (auto simp add: abs_if not_less mult_less_0_iff)
30961
541bfff659af more localisation
haftmann
parents: 30692
diff changeset
  1954
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1955
lemma abs_mult: "\<bar>a * b\<bar> = \<bar>a\<bar> * \<bar>b\<bar>"
30961
541bfff659af more localisation
haftmann
parents: 30692
diff changeset
  1956
  by (rule abs_eq_mult) auto
541bfff659af more localisation
haftmann
parents: 30692
diff changeset
  1957
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1958
lemma abs_mult_self [simp]: "\<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
  1959
  by (simp add: abs_if)
30961
541bfff659af more localisation
haftmann
parents: 30692
diff changeset
  1960
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1961
lemma abs_mult_less:
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1962
  assumes ac: "\<bar>a\<bar> < c"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1963
    and bd: "\<bar>b\<bar> < d"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1964
  shows "\<bar>a\<bar> * \<bar>b\<bar> < c * d"
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1965
proof -
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1966
  from ac have "0 < c"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1967
    by (blast intro: le_less_trans abs_ge_zero)
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1968
  with bd show ?thesis by (simp add: ac mult_strict_mono)
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1969
qed
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1970
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1971
lemma abs_less_iff: "\<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
  1972
  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
  1973
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1974
lemma abs_mult_pos: "0 \<le> x \<Longrightarrow> \<bar>y\<bar> * x = \<bar>y * x\<bar>"
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1975
  by (simp add: abs_mult)
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1976
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1977
lemma abs_diff_less_iff: "\<bar>x - a\<bar> < r \<longleftrightarrow> a - r < x \<and> x < a + r"
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
  1978
  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
  1979
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1980
lemma abs_diff_le_iff: "\<bar>x - a\<bar> \<le> r \<longleftrightarrow> a - r \<le> x \<and> x \<le> a + r"
59865
8a20dd967385 rationalised and generalised some theorems concerning abs and x^2.
paulson <lp15@cam.ac.uk>
parents: 59833
diff changeset
  1981
  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
  1982
62626
de25474ce728 Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents: 62608
diff changeset
  1983
lemma abs_add_one_gt_zero: "0 < 1 + \<bar>x\<bar>"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1984
  by (auto simp: abs_if not_less intro: zero_less_one add_strict_increasing less_trans)
62626
de25474ce728 Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents: 62608
diff changeset
  1985
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1986
end
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1987
62376
85f38d5f8807 Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents: 62366
diff changeset
  1988
subsection \<open>Dioids\<close>
85f38d5f8807 Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents: 62366
diff changeset
  1989
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1990
text \<open>
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1991
  Dioids are the alternative extensions of semirings, a semiring can
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1992
  either be a ring or a dioid but never both.
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1993
\<close>
62376
85f38d5f8807 Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents: 62366
diff changeset
  1994
85f38d5f8807 Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents: 62366
diff changeset
  1995
class dioid = semiring_1 + canonically_ordered_monoid_add
85f38d5f8807 Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents: 62366
diff changeset
  1996
begin
85f38d5f8807 Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents: 62366
diff changeset
  1997
85f38d5f8807 Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents: 62366
diff changeset
  1998
subclass ordered_semiring
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1999
  by standard (auto simp: le_iff_add distrib_left distrib_right)
62376
85f38d5f8807 Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents: 62366
diff changeset
  2000
85f38d5f8807 Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents: 62366
diff changeset
  2001
end
85f38d5f8807 Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents: 62366
diff changeset
  2002
85f38d5f8807 Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
hoelzl
parents: 62366
diff changeset
  2003
59557
ebd8ecacfba6 establish unique preferred fact names
haftmann
parents: 59555
diff changeset
  2004
hide_fact (open) comm_mult_left_mono comm_mult_strict_left_mono distrib
ebd8ecacfba6 establish unique preferred fact names
haftmann
parents: 59555
diff changeset
  2005
52435
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 51520
diff changeset
  2006
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
  2007
  code_module Rings \<rightharpoonup> (SML) Arith and (OCaml) Arith and (Haskell) Arith
33364
2bd12592c5e8 tuned code setup
haftmann
parents: 33319
diff changeset
  2008
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
  2009
end