src/HOL/Rings.thy
author haftmann
Thu, 19 Jun 2025 17:15:40 +0200
changeset 82734 89347c0cc6a3
parent 82600 f62666eea755
permissions -rw-r--r--
treat map_filter similar to list_all, list_ex, list_ex1
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
69661
a03a63b81f44 tuned proofs
haftmann
parents: 69605
diff changeset
    13
  imports Groups Set Fun
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
70145
f07b8fb99818 more document structure
haftmann
parents: 70144
diff changeset
    16
subsection \<open>Semirings and rings\<close>
f07b8fb99818 more document structure
haftmann
parents: 70144
diff changeset
    17
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
    18
class semiring = ab_semigroup_add + semigroup_mult +
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70347
diff changeset
    19
  assumes distrib_right [algebra_simps, algebra_split_simps]: "(a + b) * c = a * c + b * c"
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70347
diff changeset
    20
  assumes distrib_left [algebra_simps, algebra_split_simps]: "a * (b + c) = a * b + a * c"
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    21
begin
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    22
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    23
text \<open>For the \<open>combine_numerals\<close> simproc\<close>
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    24
lemma combine_common_factor: "a * e + (b * e + c) = (a + b) * e + c"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    25
  by (simp add: distrib_right ac_simps)
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    26
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    27
end
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    28
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
    29
class mult_zero = times + zero +
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
    30
  assumes mult_zero_left [simp]: "0 * a = 0"
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
    31
  assumes mult_zero_right [simp]: "a * 0 = 0"
58195
1fee63e0377d added various facts
haftmann
parents: 57514
diff changeset
    32
begin
1fee63e0377d added various facts
haftmann
parents: 57514
diff changeset
    33
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    34
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
    35
  by auto
1fee63e0377d added various facts
haftmann
parents: 57514
diff changeset
    36
1fee63e0377d added various facts
haftmann
parents: 57514
diff changeset
    37
end
21199
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
    38
58198
099ca49b5231 generalized
haftmann
parents: 58195
diff changeset
    39
class semiring_0 = semiring + comm_monoid_add + mult_zero
099ca49b5231 generalized
haftmann
parents: 58195
diff changeset
    40
29904
856f16a3b436 add class cancel_comm_monoid_add
huffman
parents: 29833
diff changeset
    41
class semiring_0_cancel = semiring + cancel_comm_monoid_add
25186
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
    42
begin
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    43
25186
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
    44
subclass semiring_0
28823
dcbef866c9e2 tuned unfold_locales invocation
haftmann
parents: 28559
diff changeset
    45
proof
21199
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
    46
  fix a :: 'a
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
    47
  have "0 * a + 0 * a = 0 * a + 0"
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
    48
    by (simp add: distrib_right [symmetric])
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
    49
  then show "0 * a = 0"
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
    50
    by (simp only: add_left_cancel)
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
    51
  have "a * 0 + a * 0 = a * 0 + 0"
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
    52
    by (simp add: distrib_left [symmetric])
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
    53
  then show "a * 0 = 0"
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
    54
    by (simp only: add_left_cancel)
21199
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
    55
qed
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
    56
25186
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
    57
end
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    58
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
    59
class comm_semiring = ab_semigroup_add + ab_semigroup_mult +
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
    60
  assumes distrib: "(a + b) * c = a * c + b * c"
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    61
begin
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    62
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    63
subclass semiring
28823
dcbef866c9e2 tuned unfold_locales invocation
haftmann
parents: 28559
diff changeset
    64
proof
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    65
  fix a b c :: 'a
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
    66
  show "(a + b) * c = a * c + b * c"
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
    67
    by (simp add: distrib)
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
    68
  have "a * (b + c) = (b + c) * a"
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
    69
    by (simp add: ac_simps)
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
    70
  also have "\<dots> = b * a + c * a"
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
    71
    by (simp only: distrib)
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
    72
  also have "\<dots> = a * b + a * c"
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
    73
    by (simp add: ac_simps)
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
    74
  finally show "a * (b + c) = a * b + a * c"
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
    75
    by blast
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    76
qed
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    77
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    78
end
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    79
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    80
class comm_semiring_0 = comm_semiring + comm_monoid_add + mult_zero
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    81
begin
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    82
27516
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
    83
subclass semiring_0 ..
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    84
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    85
end
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    86
29904
856f16a3b436 add class cancel_comm_monoid_add
huffman
parents: 29833
diff changeset
    87
class comm_semiring_0_cancel = comm_semiring + cancel_comm_monoid_add
25186
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
    88
begin
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
    89
27516
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
    90
subclass semiring_0_cancel ..
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
    91
28141
193c3ea0f63b instances comm_semiring_0_cancel < comm_semiring_0, comm_ring < comm_semiring_0_cancel
huffman
parents: 27651
diff changeset
    92
subclass comm_semiring_0 ..
193c3ea0f63b instances comm_semiring_0_cancel < comm_semiring_0, comm_ring < comm_semiring_0_cancel
huffman
parents: 27651
diff changeset
    93
25186
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
    94
end
21199
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
    95
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
    96
class zero_neq_one = zero + one +
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
    97
  assumes zero_neq_one [simp]: "0 \<noteq> 1"
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
    98
begin
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
    99
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
   100
lemma one_neq_zero [simp]: "1 \<noteq> 0"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   101
  by (rule not_sym) (rule zero_neq_one)
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
   102
54225
8a49a5a30284 generalized of_bool conversion
haftmann
parents: 54147
diff changeset
   103
definition of_bool :: "bool \<Rightarrow> 'a"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   104
  where "of_bool p = (if p then 1 else 0)"
54225
8a49a5a30284 generalized of_bool conversion
haftmann
parents: 54147
diff changeset
   105
8a49a5a30284 generalized of_bool conversion
haftmann
parents: 54147
diff changeset
   106
lemma of_bool_eq [simp, code]:
8a49a5a30284 generalized of_bool conversion
haftmann
parents: 54147
diff changeset
   107
  "of_bool False = 0"
8a49a5a30284 generalized of_bool conversion
haftmann
parents: 54147
diff changeset
   108
  "of_bool True = 1"
8a49a5a30284 generalized of_bool conversion
haftmann
parents: 54147
diff changeset
   109
  by (simp_all add: of_bool_def)
8a49a5a30284 generalized of_bool conversion
haftmann
parents: 54147
diff changeset
   110
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   111
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
   112
  by (simp add: of_bool_def)
8a49a5a30284 generalized of_bool conversion
haftmann
parents: 54147
diff changeset
   113
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   114
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
   115
  by (cases p) simp_all
6d0d93316daf split rules for of_bool, similar to if
haftmann
parents: 54703
diff changeset
   116
75087
f3fcc7c5a0db Avoid overaggresive splitting.
haftmann
parents: 74007
diff changeset
   117
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
   118
  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
   119
73535
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
   120
lemma of_bool_eq_0_iff [simp]:
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
   121
  \<open>of_bool P = 0 \<longleftrightarrow> \<not> P\<close>
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
   122
  by simp
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
   123
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
   124
lemma of_bool_eq_1_iff [simp]:
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
   125
  \<open>of_bool P = 1 \<longleftrightarrow> P\<close>
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
   126
  by simp
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
   127
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
   128
end
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   129
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   130
class semiring_1 = zero_neq_one + semiring_0 + monoid_mult
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66810
diff changeset
   131
begin
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66810
diff changeset
   132
70144
haftmann
parents: 70094
diff changeset
   133
lemma of_bool_conj:
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66810
diff changeset
   134
  "of_bool (P \<and> Q) = of_bool P * of_bool Q"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66810
diff changeset
   135
  by auto
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66810
diff changeset
   136
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66810
diff changeset
   137
end
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   138
71167
b4d409c65a76 Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents: 70902
diff changeset
   139
lemma lambda_zero: "(\<lambda>h::'a::mult_zero. 0) = (*) 0"
b4d409c65a76 Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents: 70902
diff changeset
   140
  by auto
b4d409c65a76 Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents: 70902
diff changeset
   141
b4d409c65a76 Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents: 70902
diff changeset
   142
lemma lambda_one: "(\<lambda>x::'a::monoid_mult. x) = (*) 1"
b4d409c65a76 Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents: 70902
diff changeset
   143
  by auto
70145
f07b8fb99818 more document structure
haftmann
parents: 70144
diff changeset
   144
f07b8fb99818 more document structure
haftmann
parents: 70144
diff changeset
   145
subsection \<open>Abstract divisibility\<close>
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   146
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   147
class dvd = times
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   148
begin
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   149
80932
261cd8722677 standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
wenzelm
parents: 79566
diff changeset
   150
definition dvd :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infix \<open>dvd\<close> 50)
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   151
  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
   152
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   153
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
   154
  unfolding dvd_def ..
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   155
68251
54a127873735 consider dvdE for automated classical proving
haftmann
parents: 67689
diff changeset
   156
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
   157
  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
   158
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   159
end
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   160
59009
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   161
context comm_monoid_mult
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   162
begin
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   163
59009
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   164
subclass dvd .
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   165
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   166
lemma dvd_refl [simp]: "a dvd a"
28559
55c003a5600a tuned default rules of (dvd)
haftmann
parents: 28141
diff changeset
   167
proof
55c003a5600a tuned default rules of (dvd)
haftmann
parents: 28141
diff changeset
   168
  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
   169
qed
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   170
62349
7c23469b5118 cleansed junk-producing interpretations for gcd/lcm on nat altogether
haftmann
parents: 62347
diff changeset
   171
lemma dvd_trans [trans]:
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   172
  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
   173
  shows "a dvd c"
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   174
proof -
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
   175
  from assms obtain v where "b = a * v"
70146
haftmann
parents: 70145
diff changeset
   176
    by auto
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
   177
  moreover from assms obtain w where "c = b * w"
70146
haftmann
parents: 70145
diff changeset
   178
    by auto
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
   179
  ultimately have "c = a * (v * w)"
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
   180
    by (simp add: mult.assoc)
28559
55c003a5600a tuned default rules of (dvd)
haftmann
parents: 28141
diff changeset
   181
  then show ?thesis ..
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   182
qed
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 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
   185
  by (auto simp add: subset_iff intro: dvd_trans)
95c6cf433c91 more theorems
haftmann
parents: 62349
diff changeset
   186
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   187
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
   188
  by (auto simp add: subset_iff intro: dvd_trans)
95c6cf433c91 more theorems
haftmann
parents: 62349
diff changeset
   189
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   190
lemma one_dvd [simp]: "1 dvd a"
70146
haftmann
parents: 70145
diff changeset
   191
  by (auto intro: dvdI)
haftmann
parents: 70145
diff changeset
   192
haftmann
parents: 70145
diff changeset
   193
lemma dvd_mult [simp]: "a dvd (b * c)" if "a dvd c"
75669
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75543
diff changeset
   194
  using that by (auto intro: mult.left_commute dvdI)
70146
haftmann
parents: 70145
diff changeset
   195
haftmann
parents: 70145
diff changeset
   196
lemma dvd_mult2 [simp]: "a dvd (b * c)" if "a dvd b"
haftmann
parents: 70145
diff changeset
   197
  using that 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
   198
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   199
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
   200
  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
   201
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   202
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
   203
  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
   204
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   205
lemma mult_dvd_mono:
30042
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29981
diff changeset
   206
  assumes "a dvd b"
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29981
diff changeset
   207
    and "c dvd d"
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   208
  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
   209
proof -
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
   210
  from \<open>a dvd b\<close> obtain b' where "b = a * b'" ..
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
   211
  moreover from \<open>c dvd d\<close> obtain d' where "d = c * d'" ..
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
   212
  ultimately have "b * d = (a * c) * (b' * d')"
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
   213
    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
   214
  then show ?thesis ..
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   215
qed
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   216
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   217
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
   218
  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
   219
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   220
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
   221
  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
   222
59009
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   223
end
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   224
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   225
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
   226
begin
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   227
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   228
subclass semiring_1 ..
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   229
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   230
lemma dvd_0_left_iff [simp]: "0 dvd a \<longleftrightarrow> a = 0"
70146
haftmann
parents: 70145
diff changeset
   231
  by auto
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   232
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   233
lemma dvd_0_right [iff]: "a dvd 0"
59009
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   234
proof
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   235
  show "0 = a * 0" by simp
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   236
qed
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   237
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   238
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
   239
  by simp
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   240
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   241
lemma dvd_add [simp]:
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   242
  assumes "a dvd b" and "a dvd c"
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   243
  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
   244
proof -
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
   245
  from \<open>a dvd b\<close> obtain b' where "b = a * b'" ..
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
   246
  moreover from \<open>a dvd c\<close> obtain c' where "c = a * c'" ..
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
   247
  ultimately have "b + c = a * (b' + c')"
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
   248
    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
   249
  then show ?thesis ..
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   250
qed
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   251
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   252
end
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   253
29904
856f16a3b436 add class cancel_comm_monoid_add
huffman
parents: 29833
diff changeset
   254
class semiring_1_cancel = semiring + cancel_comm_monoid_add
856f16a3b436 add class cancel_comm_monoid_add
huffman
parents: 29833
diff changeset
   255
  + zero_neq_one + monoid_mult
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   256
begin
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
   257
27516
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
   258
subclass semiring_0_cancel ..
25512
4134f7c782e2 using intro_locales instead of unfold_locales if appropriate
haftmann
parents: 25450
diff changeset
   259
27516
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
   260
subclass semiring_1 ..
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   261
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   262
end
21199
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
   263
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   264
class comm_semiring_1_cancel =
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   265
  comm_semiring + cancel_comm_monoid_add + zero_neq_one + comm_monoid_mult +
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70347
diff changeset
   266
  assumes right_diff_distrib' [algebra_simps, algebra_split_simps]:
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70347
diff changeset
   267
    "a * (b - c) = a * b - a * c"
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   268
begin
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   269
27516
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
   270
subclass semiring_1_cancel ..
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
   271
subclass comm_semiring_0_cancel ..
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
   272
subclass comm_semiring_1 ..
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   273
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70347
diff changeset
   274
lemma left_diff_distrib' [algebra_simps, algebra_split_simps]:
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70347
diff changeset
   275
  "(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
   276
  by (simp add: algebra_simps)
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   277
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   278
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
   279
proof -
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   280
  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
   281
  proof
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   282
    assume ?Q
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   283
    then show ?P by simp
59816
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   284
  next
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   285
    assume ?P
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   286
    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
   287
    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
   288
    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
   289
    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
   290
    then show ?Q ..
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   291
  qed
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   292
  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
   293
qed
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   294
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   295
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
   296
  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
   297
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   298
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
   299
  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
   300
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   301
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
   302
  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
   303
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   304
lemma dvd_add_right_iff:
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   305
  assumes "a dvd b"
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   306
  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
   307
proof
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   308
  assume ?P
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   309
  then obtain d where "b + c = a * d" ..
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
   310
  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
   311
  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
   312
  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
   313
  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
   314
  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
   315
  then show ?Q ..
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   316
next
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   317
  assume ?Q
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   318
  with assms show ?P by simp
59816
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   319
qed
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   320
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   321
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
   322
  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
   323
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   324
end
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   325
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   326
class ring = semiring + ab_group_add
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   327
begin
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   328
27516
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
   329
subclass semiring_0_cancel ..
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   330
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
   331
text \<open>Distribution rules\<close>
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   332
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   333
lemma minus_mult_left: "- (a * b) = - a * b"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   334
  by (rule minus_unique) (simp add: distrib_right [symmetric])
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   335
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   336
lemma minus_mult_right: "- (a * b) = a * - b"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   337
  by (rule minus_unique) (simp add: distrib_left [symmetric])
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   338
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   339
text \<open>Extract signs from products\<close>
54147
97a8ff4e4ac9 killed most "no_atp", to make Sledgehammer more complete
blanchet
parents: 52435
diff changeset
   340
lemmas mult_minus_left [simp] = minus_mult_left [symmetric]
97a8ff4e4ac9 killed most "no_atp", to make Sledgehammer more complete
blanchet
parents: 52435
diff changeset
   341
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
   342
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   343
lemma minus_mult_minus [simp]: "- a * - b = a * b"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   344
  by simp
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   345
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   346
lemma minus_mult_commute: "- a * b = a * - b"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   347
  by simp
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
   348
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70347
diff changeset
   349
lemma right_diff_distrib [algebra_simps, algebra_split_simps]:
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70347
diff changeset
   350
  "a * (b - c) = a * b - a * c"
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 54225
diff changeset
   351
  using distrib_left [of a b "-c "] by simp
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
   352
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70347
diff changeset
   353
lemma left_diff_distrib [algebra_simps, algebra_split_simps]:
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70347
diff changeset
   354
  "(a - b) * c = a * c - b * c"
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 54225
diff changeset
   355
  using distrib_right [of a "- b" c] by simp
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   356
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   357
lemmas ring_distribs = distrib_left distrib_right left_diff_distrib right_diff_distrib
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   358
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   359
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
   360
  by (simp add: algebra_simps)
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   361
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   362
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
   363
  by (simp add: algebra_simps)
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   364
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   365
end
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   366
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   367
lemmas ring_distribs = distrib_left distrib_right left_diff_distrib right_diff_distrib
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   368
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   369
class comm_ring = comm_semiring + ab_group_add
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   370
begin
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   371
27516
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
   372
subclass ring ..
28141
193c3ea0f63b instances comm_semiring_0_cancel < comm_semiring_0, comm_ring < comm_semiring_0_cancel
huffman
parents: 27651
diff changeset
   373
subclass comm_semiring_0_cancel ..
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   374
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   375
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
   376
  by (simp add: algebra_simps)
63cddfbc5a09 replace lemma realpow_two_diff with new lemma square_diff_square_factored
huffman
parents: 44346
diff changeset
   377
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   378
end
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   379
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   380
class ring_1 = ring + zero_neq_one + monoid_mult
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   381
begin
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   382
27516
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
   383
subclass semiring_1_cancel ..
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   384
75962
c530cb79ccbc avoid looping simplification for z2
haftmann
parents: 75880
diff changeset
   385
lemma of_bool_not_iff:
73535
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
   386
  \<open>of_bool (\<not> P) = 1 - of_bool P\<close>
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
   387
  by simp
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
   388
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   389
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
   390
  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
   391
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   392
end
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   393
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   394
class comm_ring_1 = comm_ring + zero_neq_one + comm_monoid_mult
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   395
begin
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   396
27516
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
   397
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
   398
subclass comm_semiring_1_cancel
70146
haftmann
parents: 70145
diff changeset
   399
  by standard (simp add: algebra_simps)
58647
fce800afeec7 more facts about abstract divisibility
haftmann
parents: 58198
diff changeset
   400
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
   401
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
   402
proof
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   403
  assume "x dvd - y"
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   404
  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
   405
  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
   406
next
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   407
  assume "x dvd y"
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   408
  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
   409
  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
   410
qed
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   411
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
   412
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
   413
proof
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   414
  assume "- x dvd y"
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   415
  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
   416
  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
   417
  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
   418
next
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   419
  assume "x dvd y"
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   420
  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
   421
  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
   422
  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
   423
qed
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   424
82518
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 80932
diff changeset
   425
lemma dvd_diff_right_iff:
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 80932
diff changeset
   426
  assumes "a dvd b"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 80932
diff changeset
   427
  shows "a dvd b - c \<longleftrightarrow> a dvd c" (is "?P \<longleftrightarrow> ?Q")
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 80932
diff changeset
   428
  using dvd_add_right_iff[of a b "-c"] assms by auto
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 80932
diff changeset
   429
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 80932
diff changeset
   430
lemma dvd_diff_left_iff: 
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 80932
diff changeset
   431
  shows "a dvd c \<Longrightarrow> a dvd b - c \<longleftrightarrow> a dvd b"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 80932
diff changeset
   432
  using dvd_add_left_iff[of a "-c" b] by auto
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 80932
diff changeset
   433
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   434
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
   435
  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
   436
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   437
end
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   438
70145
f07b8fb99818 more document structure
haftmann
parents: 70144
diff changeset
   439
f07b8fb99818 more document structure
haftmann
parents: 70144
diff changeset
   440
subsection \<open>Towards integral domains\<close>
f07b8fb99818 more document structure
haftmann
parents: 70144
diff changeset
   441
59833
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   442
class semiring_no_zero_divisors = semiring_0 +
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   443
  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
   444
begin
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   445
59833
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   446
lemma divisors_zero:
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   447
  assumes "a * b = 0"
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   448
  shows "a = 0 \<or> b = 0"
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   449
proof (rule classical)
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   450
  assume "\<not> ?thesis"
59833
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   451
  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
   452
  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
   453
  with assms show ?thesis by simp
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   454
qed
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   455
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   456
lemma mult_eq_0_iff [simp]: "a * b = 0 \<longleftrightarrow> a = 0 \<or> b = 0"
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   457
proof (cases "a = 0 \<or> b = 0")
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   458
  case False
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   459
  then have "a \<noteq> 0" and "b \<noteq> 0" by auto
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   460
    then show ?thesis using no_zero_divisors by simp
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   461
next
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   462
  case True
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   463
  then show ?thesis by auto
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   464
qed
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   465
58952
5d82cdef6c1b equivalence rules for structures without zero divisors
haftmann
parents: 58889
diff changeset
   466
end
5d82cdef6c1b equivalence rules for structures without zero divisors
haftmann
parents: 58889
diff changeset
   467
62481
b5d8e57826df tuned bootstrap order to provide type classes in a more sensible order
haftmann
parents: 62390
diff changeset
   468
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
   469
60516
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   470
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
   471
  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
   472
    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
   473
begin
5d82cdef6c1b equivalence rules for structures without zero divisors
haftmann
parents: 58889
diff changeset
   474
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   475
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
   476
  by simp
56217
dc429a5b13c4 Some rationalisation of basic lemmas
paulson <lp15@cam.ac.uk>
parents: 55912
diff changeset
   477
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   478
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
   479
  by simp
56217
dc429a5b13c4 Some rationalisation of basic lemmas
paulson <lp15@cam.ac.uk>
parents: 55912
diff changeset
   480
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   481
end
22990
775e9de3db48 added classes ring_no_zero_divisors and dom (non-commutative version of idom);
huffman
parents: 22987
diff changeset
   482
60516
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   483
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
   484
begin
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   485
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   486
subclass semiring_no_zero_divisors_cancel
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   487
proof
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   488
  fix a b c
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   489
  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
   490
    by (simp add: algebra_simps)
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   491
  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
   492
    by auto
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   493
  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
   494
  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
   495
    by (simp add: algebra_simps)
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   496
  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
   497
    by auto
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   498
  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
   499
qed
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   500
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   501
end
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   502
23544
4b4165cb3e0d rename class dom to ring_1_no_zero_divisors
huffman
parents: 23527
diff changeset
   503
class ring_1_no_zero_divisors = ring_1 + ring_no_zero_divisors
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
   504
begin
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
   505
62481
b5d8e57826df tuned bootstrap order to provide type classes in a more sensible order
haftmann
parents: 62390
diff changeset
   506
subclass semiring_1_no_zero_divisors ..
b5d8e57826df tuned bootstrap order to provide type classes in a more sensible order
haftmann
parents: 62390
diff changeset
   507
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   508
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
   509
proof -
9207505d1ee5 move lemma real_mult_is_one to Rings.thy, renamed to square_eq_1_iff
huffman
parents: 36719
diff changeset
   510
  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
   511
    by (simp add: algebra_simps)
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   512
  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
   513
    by simp
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   514
  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
   515
    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
   516
qed
9207505d1ee5 move lemma real_mult_is_one to Rings.thy, renamed to square_eq_1_iff
huffman
parents: 36719
diff changeset
   517
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   518
lemma mult_cancel_right1 [simp]: "c = b * c \<longleftrightarrow> c = 0 \<or> b = 1"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   519
  using mult_cancel_right [of 1 c b] by auto
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
   520
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   521
lemma mult_cancel_right2 [simp]: "a * c = c \<longleftrightarrow> c = 0 \<or> a = 1"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   522
  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
   523
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   524
lemma mult_cancel_left1 [simp]: "c = c * b \<longleftrightarrow> c = 0 \<or> b = 1"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   525
  using mult_cancel_left [of c 1 b] by force
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
   526
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   527
lemma mult_cancel_left2 [simp]: "c * a = c \<longleftrightarrow> c = 0 \<or> a = 1"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   528
  using mult_cancel_left [of c a 1] by simp
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
   529
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
   530
end
22990
775e9de3db48 added classes ring_no_zero_divisors and dom (non-commutative version of idom);
huffman
parents: 22987
diff changeset
   531
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
   532
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
   533
begin
b5d8e57826df tuned bootstrap order to provide type classes in a more sensible order
haftmann
parents: 62390
diff changeset
   534
b5d8e57826df tuned bootstrap order to provide type classes in a more sensible order
haftmann
parents: 62390
diff changeset
   535
subclass semiring_1_no_zero_divisors ..
b5d8e57826df tuned bootstrap order to provide type classes in a more sensible order
haftmann
parents: 62390
diff changeset
   536
b5d8e57826df tuned bootstrap order to provide type classes in a more sensible order
haftmann
parents: 62390
diff changeset
   537
end
59833
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   538
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   539
class idom = comm_ring_1 + semiring_no_zero_divisors
25186
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
   540
begin
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   541
59833
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   542
subclass semidom ..
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   543
27516
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
   544
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
   545
70146
haftmann
parents: 70145
diff changeset
   546
lemma dvd_mult_cancel_right [simp]:
haftmann
parents: 70145
diff changeset
   547
  "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
   548
proof -
7d0ed261b712 generalize int_dvd_cancel_factor simproc to idom class
huffman
parents: 29949
diff changeset
   549
  have "a * c dvd b * c \<longleftrightarrow> (\<exists>k. b * c = (a * k) * c)"
70146
haftmann
parents: 70145
diff changeset
   550
    by (auto simp add: ac_simps)
29981
7d0ed261b712 generalize int_dvd_cancel_factor simproc to idom class
huffman
parents: 29949
diff changeset
   551
  also have "(\<exists>k. b * c = (a * k) * c) \<longleftrightarrow> c = 0 \<or> a dvd b"
70146
haftmann
parents: 70145
diff changeset
   552
    by auto
29981
7d0ed261b712 generalize int_dvd_cancel_factor simproc to idom class
huffman
parents: 29949
diff changeset
   553
  finally show ?thesis .
7d0ed261b712 generalize int_dvd_cancel_factor simproc to idom class
huffman
parents: 29949
diff changeset
   554
qed
7d0ed261b712 generalize int_dvd_cancel_factor simproc to idom class
huffman
parents: 29949
diff changeset
   555
70146
haftmann
parents: 70145
diff changeset
   556
lemma dvd_mult_cancel_left [simp]:
haftmann
parents: 70145
diff changeset
   557
  "c * a dvd c * b \<longleftrightarrow> c = 0 \<or> a dvd b"
haftmann
parents: 70145
diff changeset
   558
  using dvd_mult_cancel_right [of a c b] by (simp add: ac_simps)
29981
7d0ed261b712 generalize int_dvd_cancel_factor simproc to idom class
huffman
parents: 29949
diff changeset
   559
60516
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   560
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
   561
proof
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   562
  assume "a * a = b * b"
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   563
  then have "(a - b) * (a + b) = 0"
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   564
    by (simp add: algebra_simps)
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   565
  then show "a = b \<or> a = - b"
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   566
    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
   567
next
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   568
  assume "a = b \<or> a = - b"
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   569
  then show "a * a = b * b" by auto
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   570
qed
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   571
75880
714fad33252e more thorough split rules for div and mod on numerals, tuned split rules setup
haftmann
parents: 75875
diff changeset
   572
lemma inj_mult_left [simp]: \<open>inj ((*) a) \<longleftrightarrow> a \<noteq> 0\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
714fad33252e more thorough split rules for div and mod on numerals, tuned split rules setup
haftmann
parents: 75875
diff changeset
   573
proof
714fad33252e more thorough split rules for div and mod on numerals, tuned split rules setup
haftmann
parents: 75875
diff changeset
   574
  assume ?P
714fad33252e more thorough split rules for div and mod on numerals, tuned split rules setup
haftmann
parents: 75875
diff changeset
   575
  show ?Q
714fad33252e more thorough split rules for div and mod on numerals, tuned split rules setup
haftmann
parents: 75875
diff changeset
   576
  proof
714fad33252e more thorough split rules for div and mod on numerals, tuned split rules setup
haftmann
parents: 75875
diff changeset
   577
    assume \<open>a = 0\<close>
714fad33252e more thorough split rules for div and mod on numerals, tuned split rules setup
haftmann
parents: 75875
diff changeset
   578
    with \<open>?P\<close> have "inj ((*) 0)"
714fad33252e more thorough split rules for div and mod on numerals, tuned split rules setup
haftmann
parents: 75875
diff changeset
   579
      by simp
714fad33252e more thorough split rules for div and mod on numerals, tuned split rules setup
haftmann
parents: 75875
diff changeset
   580
    moreover have "0 * 0 = 0 * 1"
714fad33252e more thorough split rules for div and mod on numerals, tuned split rules setup
haftmann
parents: 75875
diff changeset
   581
      by simp
714fad33252e more thorough split rules for div and mod on numerals, tuned split rules setup
haftmann
parents: 75875
diff changeset
   582
    ultimately have "0 = 1"
714fad33252e more thorough split rules for div and mod on numerals, tuned split rules setup
haftmann
parents: 75875
diff changeset
   583
      by (rule injD)
714fad33252e more thorough split rules for div and mod on numerals, tuned split rules setup
haftmann
parents: 75875
diff changeset
   584
    then show False
714fad33252e more thorough split rules for div and mod on numerals, tuned split rules setup
haftmann
parents: 75875
diff changeset
   585
      by simp
714fad33252e more thorough split rules for div and mod on numerals, tuned split rules setup
haftmann
parents: 75875
diff changeset
   586
  qed
714fad33252e more thorough split rules for div and mod on numerals, tuned split rules setup
haftmann
parents: 75875
diff changeset
   587
next
714fad33252e more thorough split rules for div and mod on numerals, tuned split rules setup
haftmann
parents: 75875
diff changeset
   588
  assume ?Q then show ?P
714fad33252e more thorough split rules for div and mod on numerals, tuned split rules setup
haftmann
parents: 75875
diff changeset
   589
    by (auto intro: injI)
714fad33252e more thorough split rules for div and mod on numerals, tuned split rules setup
haftmann
parents: 75875
diff changeset
   590
qed
714fad33252e more thorough split rules for div and mod on numerals, tuned split rules setup
haftmann
parents: 75875
diff changeset
   591
25186
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
   592
end
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   593
64290
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   594
class idom_abs_sgn = idom + abs + sgn +
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   595
  assumes sgn_mult_abs: "sgn a * \<bar>a\<bar> = a"
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   596
    and sgn_sgn [simp]: "sgn (sgn a) = sgn a"
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   597
    and abs_abs [simp]: "\<bar>\<bar>a\<bar>\<bar> = \<bar>a\<bar>"
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   598
    and abs_0 [simp]: "\<bar>0\<bar> = 0"
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   599
    and sgn_0 [simp]: "sgn 0 = 0"
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   600
    and sgn_1 [simp]: "sgn 1 = 1"
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   601
    and sgn_minus_1: "sgn (- 1) = - 1"
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   602
    and sgn_mult: "sgn (a * b) = sgn a * sgn b"
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   603
begin
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   604
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   605
lemma sgn_eq_0_iff:
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   606
  "sgn a = 0 \<longleftrightarrow> a = 0"
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   607
proof -
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   608
  { assume "sgn a = 0"
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   609
    then have "sgn a * \<bar>a\<bar> = 0"
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   610
      by simp
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   611
    then have "a = 0"
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   612
      by (simp add: sgn_mult_abs)
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   613
  } then show ?thesis
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   614
    by auto
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   615
qed
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   616
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   617
lemma abs_eq_0_iff:
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   618
  "\<bar>a\<bar> = 0 \<longleftrightarrow> a = 0"
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   619
proof -
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   620
  { assume "\<bar>a\<bar> = 0"
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   621
    then have "sgn a * \<bar>a\<bar> = 0"
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   622
      by simp
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   623
    then have "a = 0"
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   624
      by (simp add: sgn_mult_abs)
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   625
  } then show ?thesis
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   626
    by auto
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   627
qed
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   628
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   629
lemma abs_mult_sgn:
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   630
  "\<bar>a\<bar> * sgn a = a"
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   631
  using sgn_mult_abs [of a] by (simp add: ac_simps)
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   632
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   633
lemma abs_1 [simp]:
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   634
  "\<bar>1\<bar> = 1"
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   635
  using sgn_mult_abs [of 1] by simp
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   636
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   637
lemma sgn_abs [simp]:
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   638
  "\<bar>sgn a\<bar> = of_bool (a \<noteq> 0)"
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   639
  using sgn_mult_abs [of "sgn a"] mult_cancel_left [of "sgn a" "\<bar>sgn a\<bar>" 1]
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   640
  by (auto simp add: sgn_eq_0_iff)
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   641
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   642
lemma abs_sgn [simp]:
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   643
  "sgn \<bar>a\<bar> = of_bool (a \<noteq> 0)"
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   644
  using sgn_mult_abs [of "\<bar>a\<bar>"] mult_cancel_right [of "sgn \<bar>a\<bar>" "\<bar>a\<bar>" 1]
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   645
  by (auto simp add: abs_eq_0_iff)
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   646
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   647
lemma abs_mult:
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   648
  "\<bar>a * b\<bar> = \<bar>a\<bar> * \<bar>b\<bar>"
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   649
proof (cases "a = 0 \<or> b = 0")
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   650
  case True
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   651
  then show ?thesis
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   652
    by auto
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   653
next
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   654
  case False
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   655
  then have *: "sgn (a * b) \<noteq> 0"
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   656
    by (simp add: sgn_eq_0_iff)
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   657
  from abs_mult_sgn [of "a * b"] abs_mult_sgn [of a] abs_mult_sgn [of b]
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   658
  have "\<bar>a * b\<bar> * sgn (a * b) = \<bar>a\<bar> * sgn a * \<bar>b\<bar> * sgn b"
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   659
    by (simp add: ac_simps)
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   660
  then have "\<bar>a * b\<bar> * sgn (a * b) = \<bar>a\<bar> * \<bar>b\<bar> * sgn (a * b)"
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   661
    by (simp add: sgn_mult ac_simps)
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   662
  with * show ?thesis
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   663
    by simp
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   664
qed
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   665
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   666
lemma sgn_minus [simp]:
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   667
  "sgn (- a) = - sgn a"
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   668
proof -
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   669
  from sgn_minus_1 have "sgn (- 1 * a) = - 1 * sgn a"
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   670
    by (simp only: sgn_mult)
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   671
  then show ?thesis
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   672
    by simp
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   673
qed
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   674
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   675
lemma abs_minus [simp]:
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   676
  "\<bar>- a\<bar> = \<bar>a\<bar>"
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   677
proof -
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   678
  have [simp]: "\<bar>- 1\<bar> = 1"
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   679
    using sgn_mult_abs [of "- 1"] by simp
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   680
  then have "\<bar>- 1 * a\<bar> = 1 * \<bar>a\<bar>"
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   681
    by (simp only: abs_mult)
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   682
  then show ?thesis
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   683
    by simp
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   684
qed
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   685
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   686
end
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
   687
70145
f07b8fb99818 more document structure
haftmann
parents: 70144
diff changeset
   688
f07b8fb99818 more document structure
haftmann
parents: 70144
diff changeset
   689
subsection \<open>(Partial) Division\<close>
63950
cdc1e59aa513 syntactic type class for operation mod named after mod;
haftmann
parents: 63947
diff changeset
   690
60353
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   691
class divide =
80932
261cd8722677 standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
wenzelm
parents: 79566
diff changeset
   692
  fixes divide :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  (infixl \<open>div\<close> 70)
60353
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   693
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68253
diff changeset
   694
setup \<open>Sign.add_const_constraint (\<^const_name>\<open>divide\<close>, SOME \<^typ>\<open>'a \<Rightarrow> 'a \<Rightarrow> 'a\<close>)\<close>
60353
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   695
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   696
context semiring
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   697
begin
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   698
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70347
diff changeset
   699
lemma [field_simps, field_split_simps]:
60429
d3d1e185cd63 uniform _ div _ as infix syntax for ring division
haftmann
parents: 60353
diff changeset
   700
  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
   701
    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
   702
  by (rule distrib_left distrib_right)+
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   703
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   704
end
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   705
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   706
context ring
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   707
begin
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   708
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70347
diff changeset
   709
lemma [field_simps, field_split_simps]:
60429
d3d1e185cd63 uniform _ div _ as infix syntax for ring division
haftmann
parents: 60353
diff changeset
   710
  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
   711
    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
   712
  by (rule left_diff_distrib right_diff_distrib)+
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   713
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   714
end
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   715
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68253
diff changeset
   716
setup \<open>Sign.add_const_constraint (\<^const_name>\<open>divide\<close>, SOME \<^typ>\<open>'a::divide \<Rightarrow> 'a \<Rightarrow> 'a\<close>)\<close>
60353
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   717
79541
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
   718
class divide_trivial = zero + one + divide +
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
   719
  assumes div_by_0 [simp]: \<open>a div 0 = 0\<close>
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
   720
    and div_by_1 [simp]: \<open>a div 1 = a\<close>
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
   721
    and div_0 [simp]: \<open>0 div a = 0\<close>
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
   722
  
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
   723
63950
cdc1e59aa513 syntactic type class for operation mod named after mod;
haftmann
parents: 63947
diff changeset
   724
text \<open>Algebraic classes with division\<close>
cdc1e59aa513 syntactic type class for operation mod named after mod;
haftmann
parents: 63947
diff changeset
   725
  
60353
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   726
class semidom_divide = semidom + divide +
79541
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
   727
  assumes nonzero_mult_div_cancel_right [simp]: \<open>b \<noteq> 0 \<Longrightarrow> (a * b) div b = a\<close>
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
   728
  assumes semidom_div_by_0: \<open>a div 0 = 0\<close>
60353
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   729
begin
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   730
79541
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
   731
lemma nonzero_mult_div_cancel_left [simp]: \<open>a \<noteq> 0 \<Longrightarrow> (a * b) div a = b\<close>
64240
eabf80376aab more standardized names
haftmann
parents: 64239
diff changeset
   732
  using nonzero_mult_div_cancel_right [of a b] by (simp add: ac_simps)
60353
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   733
79541
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
   734
subclass divide_trivial
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
   735
proof
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
   736
  show [simp]: \<open>a div 0 = 0\<close> for a
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
   737
    by (fact semidom_div_by_0)
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
   738
  show \<open>a div 1 = a\<close> for a
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
   739
    using nonzero_mult_div_cancel_right [of 1 a] by simp
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
   740
  show \<open>0 div a = 0\<close> for a
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
   741
    using nonzero_mult_div_cancel_right [of a 0] by (cases \<open>a = 0\<close>) simp_all
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
   742
qed
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
   743
60516
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   744
subclass semiring_no_zero_divisors_cancel
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   745
proof
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   746
  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
   747
  proof (cases "c = 0")
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   748
    case True
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   749
    then show ?thesis by simp
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   750
  next
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   751
    case False
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
   752
    have "a = b" if "a * c = b * c"
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
   753
    proof -
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
   754
      from that have "a * c div c = b * c div c"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   755
        by simp
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
   756
      with False show ?thesis
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   757
        by simp
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
   758
    qed
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   759
    then show ?thesis by auto
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   760
  qed
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   761
  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
   762
    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
   763
qed
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   764
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   765
lemma div_self [simp]: "a \<noteq> 0 \<Longrightarrow> a div a = 1"
64240
eabf80376aab more standardized names
haftmann
parents: 64239
diff changeset
   766
  using nonzero_mult_div_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
   767
64591
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
   768
lemma dvd_div_eq_0_iff:
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
   769
  assumes "b dvd a"
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
   770
  shows "a div b = 0 \<longleftrightarrow> a = 0"
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
   771
  using assms by (elim dvdE, cases "b = 0") simp_all  
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
   772
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
   773
lemma dvd_div_eq_cancel:
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
   774
  "a div c = b div c \<Longrightarrow> c dvd a \<Longrightarrow> c dvd b \<Longrightarrow> a = b"
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
   775
  by (elim dvdE, cases "c = 0") simp_all
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
   776
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
   777
lemma dvd_div_eq_iff:
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
   778
  "c dvd a \<Longrightarrow> c dvd b \<Longrightarrow> a div c = b div c \<longleftrightarrow> a = b"
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
   779
  by (elim dvdE, cases "c = 0") simp_all
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
   780
69661
a03a63b81f44 tuned proofs
haftmann
parents: 69605
diff changeset
   781
lemma inj_on_mult:
a03a63b81f44 tuned proofs
haftmann
parents: 69605
diff changeset
   782
  "inj_on ((*) a) A" if "a \<noteq> 0"
a03a63b81f44 tuned proofs
haftmann
parents: 69605
diff changeset
   783
proof (rule inj_onI)
a03a63b81f44 tuned proofs
haftmann
parents: 69605
diff changeset
   784
  fix b c
a03a63b81f44 tuned proofs
haftmann
parents: 69605
diff changeset
   785
  assume "a * b = a * c"
a03a63b81f44 tuned proofs
haftmann
parents: 69605
diff changeset
   786
  then have "a * b div a = a * c div a"
a03a63b81f44 tuned proofs
haftmann
parents: 69605
diff changeset
   787
    by (simp only:)
a03a63b81f44 tuned proofs
haftmann
parents: 69605
diff changeset
   788
  with that show "b = c"
a03a63b81f44 tuned proofs
haftmann
parents: 69605
diff changeset
   789
    by simp
a03a63b81f44 tuned proofs
haftmann
parents: 69605
diff changeset
   790
qed
a03a63b81f44 tuned proofs
haftmann
parents: 69605
diff changeset
   791
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   792
end
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   793
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   794
class idom_divide = idom + semidom_divide
64591
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
   795
begin
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
   796
64592
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
   797
lemma dvd_neg_div:
64591
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
   798
  assumes "b dvd a"
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
   799
  shows "- a div b = - (a div b)"
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
   800
proof (cases "b = 0")
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
   801
  case True
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
   802
  then show ?thesis by simp
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
   803
next
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
   804
  case False
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
   805
  from assms obtain c where "a = b * c" ..
64592
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
   806
  then have "- a div b = (b * - c) div b"
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
   807
    by simp
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
   808
  from False also have "\<dots> = - c"
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
   809
    by (rule nonzero_mult_div_cancel_left)  
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
   810
  with False \<open>a = b * c\<close> show ?thesis
64591
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
   811
    by simp
64592
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
   812
qed
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
   813
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
   814
lemma dvd_div_neg:
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
   815
  assumes "b dvd a"
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
   816
  shows "a div - b = - (a div b)"
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
   817
proof (cases "b = 0")
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
   818
  case True
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
   819
  then show ?thesis by simp
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
   820
next
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
   821
  case False
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
   822
  then have "- b \<noteq> 0"
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
   823
    by simp
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
   824
  from assms obtain c where "a = b * c" ..
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
   825
  then have "a div - b = (- b * - c) div - b"
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
   826
    by simp
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
   827
  from \<open>- b \<noteq> 0\<close> also have "\<dots> = - c"
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
   828
    by (rule nonzero_mult_div_cancel_left)  
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
   829
  with False \<open>a = b * c\<close> show ?thesis
64591
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
   830
    by simp
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
   831
qed
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
   832
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
   833
end
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   834
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   835
class algebraic_semidom = semidom_divide
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   836
begin
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   837
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   838
text \<open>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68253
diff changeset
   839
  Class \<^class>\<open>algebraic_semidom\<close> enriches a integral domain
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   840
  by notions from algebra, like units in a ring.
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   841
  It is a separate class to avoid spoiling fields with notions
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   842
  which are degenerated there.
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   843
\<close>
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   844
60690
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   845
lemma dvd_times_left_cancel_iff [simp]:
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   846
  assumes "a \<noteq> 0"
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
   847
  shows "a * b dvd a * c \<longleftrightarrow> b dvd c"
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
   848
    (is "?lhs \<longleftrightarrow> ?rhs")
60690
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   849
proof
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
   850
  assume ?lhs
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   851
  then obtain d where "a * c = a * b * d" ..
60690
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   852
  with assms have "c = b * d" by (simp add: ac_simps)
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
   853
  then show ?rhs ..
60690
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   854
next
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
   855
  assume ?rhs
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   856
  then obtain d where "c = b * d" ..
60690
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   857
  then have "a * c = a * b * d" by (simp add: ac_simps)
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
   858
  then show ?lhs ..
60690
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   859
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
   860
60690
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   861
lemma dvd_times_right_cancel_iff [simp]:
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   862
  assumes "a \<noteq> 0"
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
   863
  shows "b * a dvd c * a \<longleftrightarrow> b dvd c"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   864
  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
   865
60690
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   866
lemma div_dvd_iff_mult:
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   867
  assumes "b \<noteq> 0" and "b dvd a"
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   868
  shows "a div b dvd c \<longleftrightarrow> a dvd c * b"
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   869
proof -
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   870
  from \<open>b dvd a\<close> obtain d where "a = b * d" ..
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   871
  with \<open>b \<noteq> 0\<close> show ?thesis by (simp add: ac_simps)
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   872
qed
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   873
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   874
lemma dvd_div_iff_mult:
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   875
  assumes "c \<noteq> 0" and "c dvd b"
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   876
  shows "a dvd b div c \<longleftrightarrow> a * c dvd b"
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   877
proof -
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   878
  from \<open>c dvd b\<close> obtain d where "b = c * d" ..
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   879
  with \<open>c \<noteq> 0\<close> show ?thesis by (simp add: mult.commute [of a])
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   880
qed
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   881
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   882
lemma div_dvd_div [simp]:
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   883
  assumes "a dvd b" and "a dvd c"
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   884
  shows "b div a dvd c div a \<longleftrightarrow> b dvd c"
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   885
proof (cases "a = 0")
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   886
  case True
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   887
  with assms show ?thesis by simp
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   888
next
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   889
  case False
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   890
  moreover from assms obtain k l where "b = a * k" and "c = a * l"
70146
haftmann
parents: 70145
diff changeset
   891
    by blast
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   892
  ultimately show ?thesis by simp
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   893
qed
60353
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   894
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   895
lemma div_add [simp]:
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   896
  assumes "c dvd a" and "c dvd b"
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   897
  shows "(a + b) div c = a div c + b div c"
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   898
proof (cases "c = 0")
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   899
  case True
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   900
  then show ?thesis by simp
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   901
next
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   902
  case False
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   903
  moreover from assms obtain k l where "a = c * k" and "b = c * l"
70146
haftmann
parents: 70145
diff changeset
   904
    by blast
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   905
  moreover have "c * k + c * l = c * (k + l)"
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   906
    by (simp add: algebra_simps)
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   907
  ultimately show ?thesis
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   908
    by simp
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   909
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
   910
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   911
lemma div_mult_div_if_dvd:
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   912
  assumes "b dvd a" and "d dvd c"
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   913
  shows "(a div b) * (c div d) = (a * c) div (b * d)"
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   914
proof (cases "b = 0 \<or> c = 0")
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   915
  case True
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   916
  with assms show ?thesis by auto
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   917
next
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   918
  case False
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   919
  moreover from assms obtain k l where "a = b * k" and "c = d * l"
70146
haftmann
parents: 70145
diff changeset
   920
    by blast
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   921
  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
   922
    by (simp add: ac_simps)
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   923
  ultimately show ?thesis by simp
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   924
qed
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   925
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   926
lemma dvd_div_eq_mult:
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   927
  assumes "a \<noteq> 0" and "a dvd b"
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   928
  shows "b div a = c \<longleftrightarrow> b = c * a"
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
   929
    (is "?lhs \<longleftrightarrow> ?rhs")
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   930
proof
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
   931
  assume ?rhs
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
   932
  then show ?lhs by (simp add: assms)
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   933
next
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
   934
  assume ?lhs
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   935
  then have "b div a * a = c * a" by simp
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   936
  moreover from assms have "b div a * a = b"
70146
haftmann
parents: 70145
diff changeset
   937
    by (auto simp add: ac_simps)
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
   938
  ultimately show ?rhs by simp
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   939
qed
60688
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
   940
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   941
lemma dvd_div_mult_self [simp]: "a dvd b \<Longrightarrow> b div a * a = b"
70146
haftmann
parents: 70145
diff changeset
   942
  by (cases "a = 0") (auto 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
   943
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   944
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
   945
  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
   946
60517
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   947
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
   948
  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
   949
  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
   950
proof (cases "c = 0")
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   951
  case True
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   952
  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
   953
next
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   954
  case False
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   955
  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
   956
  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
   957
    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
   958
  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
   959
qed
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   960
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   961
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
   962
  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
   963
60570
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   964
lemma dvd_div_mult2_eq:
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   965
  assumes "b * c dvd a"
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   966
  shows "a div (b * c) = a div b div c"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   967
proof -
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   968
  from assms obtain k where "a = b * c * k" ..
60570
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   969
  then show ?thesis
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   970
    by (cases "b = 0 \<or> c = 0") (auto, simp add: ac_simps)
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   971
qed
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   972
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   973
lemma dvd_div_div_eq_mult:
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   974
  assumes "a \<noteq> 0" "c \<noteq> 0" and "a dvd b" "c dvd d"
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
   975
  shows "b div a = d div c \<longleftrightarrow> b * c = a * d"
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
   976
    (is "?lhs \<longleftrightarrow> ?rhs")
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   977
proof -
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   978
  from assms have "a * c \<noteq> 0" by simp
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
   979
  then have "?lhs \<longleftrightarrow> b div a * (a * c) = d div c * (a * c)"
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   980
    by simp
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   981
  also have "\<dots> \<longleftrightarrow> (a * (b div a)) * c = (c * (d div c)) * a"
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   982
    by (simp add: ac_simps)
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   983
  also have "\<dots> \<longleftrightarrow> (a * b div a) * c = (c * d div c) * a"
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   984
    using assms by (simp add: div_mult_swap)
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
   985
  also have "\<dots> \<longleftrightarrow> ?rhs"
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   986
    using assms by (simp add: ac_simps)
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   987
  finally show ?thesis .
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   988
qed
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60758
diff changeset
   989
63359
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63325
diff changeset
   990
lemma dvd_mult_imp_div:
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63325
diff changeset
   991
  assumes "a * c dvd b"
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63325
diff changeset
   992
  shows "a dvd b div c"
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63325
diff changeset
   993
proof (cases "c = 0")
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63325
diff changeset
   994
  case True then show ?thesis by simp
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63325
diff changeset
   995
next
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63325
diff changeset
   996
  case False
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63325
diff changeset
   997
  from \<open>a * c dvd b\<close> obtain d where "b = a * c * d" ..
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
   998
  with False show ?thesis
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
   999
    by (simp add: mult.commute [of a] mult.assoc)
63359
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63325
diff changeset
  1000
qed
99b51ba8da1c More lemmas on Gcd/Lcm
Manuel Eberl <eberlm@in.tum.de>
parents: 63325
diff changeset
  1001
64592
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
  1002
lemma div_div_eq_right:
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
  1003
  assumes "c dvd b" "b dvd a"
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
  1004
  shows   "a div (b div c) = a div b * c"
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
  1005
proof (cases "c = 0 \<or> b = 0")
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
  1006
  case True
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
  1007
  then show ?thesis
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
  1008
    by auto
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
  1009
next
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
  1010
  case False
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
  1011
  from assms obtain r s where "b = c * r" and "a = c * r * s"
70146
haftmann
parents: 70145
diff changeset
  1012
    by blast
64592
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
  1013
  moreover with False have "r \<noteq> 0"
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
  1014
    by auto
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
  1015
  ultimately show ?thesis using False
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
  1016
    by simp (simp add: mult.commute [of _ r] mult.assoc mult.commute [of c])
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
  1017
qed
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
  1018
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
  1019
lemma div_div_div_same:
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
  1020
  assumes "d dvd b" "b dvd a"
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
  1021
  shows   "(a div d) div (b div d) = a div b"
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
  1022
proof (cases "b = 0 \<or> d = 0")
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
  1023
  case True
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
  1024
  with assms show ?thesis
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
  1025
    by auto
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
  1026
next
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
  1027
  case False
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
  1028
  from assms obtain r s
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
  1029
    where "a = d * r * s" and "b = d * r"
70146
haftmann
parents: 70145
diff changeset
  1030
    by blast
64592
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
  1031
  with False show ?thesis
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
  1032
    by simp (simp add: ac_simps)
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
  1033
qed
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64591
diff changeset
  1034
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
  1035
60517
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
  1036
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
  1037
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
  1038
abbreviation is_unit :: "'a \<Rightarrow> bool"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1039
  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
  1040
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1041
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
  1042
  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
  1043
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1044
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
  1045
  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
  1046
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
  1047
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
  1048
  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
  1049
  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
  1050
proof -
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
  1051
  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
  1052
  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
  1053
  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
  1054
  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
  1055
qed
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
  1056
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1057
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
  1058
  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
  1059
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
  1060
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
  1061
  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
  1062
  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
  1063
proof -
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
  1064
  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
  1065
  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
  1066
qed
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
  1067
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
  1068
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
  1069
  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
  1070
  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
  1071
    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
  1072
    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
  1073
proof (rule that)
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62626
diff changeset
  1074
  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
  1075
  then show "1 div a = b" by simp
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1076
  from assms b_def show "is_unit b" by simp
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1077
  with assms show "a \<noteq> 0" and "b \<noteq> 0" by auto
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1078
  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
  1079
  then have "1 = a * b" ..
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
  1080
  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
  1081
  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
  1082
  then obtain d where "c = a * d" ..
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
  1083
  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
  1084
    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
  1085
qed
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
  1086
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1087
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
  1088
  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
  1089
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1090
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
  1091
  by (auto dest: dvd_mult_left dvd_mult_right)
95c6cf433c91 more theorems
haftmann
parents: 62349
diff changeset
  1092
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1093
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
  1094
  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
  1095
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
  1096
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
  1097
  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
  1098
  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
  1099
proof
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
  1100
  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
  1101
  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
  1102
    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
  1103
next
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
  1104
  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
  1105
  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
  1106
  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
  1107
    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
  1108
  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
  1109
qed
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
  1110
63924
f91766530e13 more generic algebraic lemmas
haftmann
parents: 63680
diff changeset
  1111
lemma mult_unit_dvd_iff': "is_unit a \<Longrightarrow> (a * b) dvd c \<longleftrightarrow> b dvd c"
f91766530e13 more generic algebraic lemmas
haftmann
parents: 63680
diff changeset
  1112
  using mult_unit_dvd_iff [of a b c] by (simp add: ac_simps)
f91766530e13 more generic algebraic lemmas
haftmann
parents: 63680
diff changeset
  1113
60517
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
  1114
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
  1115
  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
  1116
  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
  1117
proof
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
  1118
  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
  1119
  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
  1120
    by (subst mult_assoc [symmetric]) simp
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1121
  also from assms have "b * (1 div b) = 1"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1122
    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
  1123
  finally have "c * b dvd c" by simp
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
  1124
  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
  1125
next
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
  1126
  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
  1127
  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
  1128
qed
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
  1129
63924
f91766530e13 more generic algebraic lemmas
haftmann
parents: 63680
diff changeset
  1130
lemma dvd_mult_unit_iff': "is_unit b \<Longrightarrow> a dvd b * c \<longleftrightarrow> a dvd c"
f91766530e13 more generic algebraic lemmas
haftmann
parents: 63680
diff changeset
  1131
  using dvd_mult_unit_iff [of b a c] by (simp add: ac_simps)
f91766530e13 more generic algebraic lemmas
haftmann
parents: 63680
diff changeset
  1132
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1133
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
  1134
  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
  1135
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1136
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
  1137
  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
  1138
63924
f91766530e13 more generic algebraic lemmas
haftmann
parents: 63680
diff changeset
  1139
lemmas unit_dvd_iff = mult_unit_dvd_iff mult_unit_dvd_iff'
f91766530e13 more generic algebraic lemmas
haftmann
parents: 63680
diff changeset
  1140
  dvd_mult_unit_iff dvd_mult_unit_iff' 
f91766530e13 more generic algebraic lemmas
haftmann
parents: 63680
diff changeset
  1141
  div_unit_dvd_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
  1142
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1143
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
  1144
  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
  1145
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1146
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
  1147
  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
  1148
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1149
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
  1150
  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
  1151
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1152
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
  1153
  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
  1154
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1155
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
  1156
  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
  1157
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1158
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
  1159
  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
  1160
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1161
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
  1162
  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
  1163
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1164
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
  1165
  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
  1166
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1167
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
  1168
  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
  1169
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
  1170
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
  1171
  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
  1172
  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
  1173
proof -
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
  1174
  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
  1175
  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
  1176
    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
  1177
  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
  1178
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
  1179
60570
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
  1180
lemma is_unit_div_mult2_eq:
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
  1181
  assumes "is_unit b" and "is_unit c"
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
  1182
  shows "a div (b * c) = a div b div c"
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
  1183
proof -
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1184
  from assms have "is_unit (b * c)"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1185
    by (simp add: unit_prod)
60570
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
  1186
  then have "b * c dvd a"
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
  1187
    by (rule unit_imp_dvd)
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
  1188
  then show ?thesis
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
  1189
    by (rule dvd_div_mult2_eq)
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
  1190
qed
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
  1191
64240
eabf80376aab more standardized names
haftmann
parents: 64239
diff changeset
  1192
lemma is_unit_div_mult_cancel_left:
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1193
  assumes "a \<noteq> 0" and "is_unit b"
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1194
  shows "a div (a * b) = 1 div b"
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1195
proof -
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1196
  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
  1197
    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
  1198
  with assms show ?thesis by simp
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1199
qed
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1200
64240
eabf80376aab more standardized names
haftmann
parents: 64239
diff changeset
  1201
lemma is_unit_div_mult_cancel_right:
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1202
  assumes "a \<noteq> 0" and "is_unit b"
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1203
  shows "a div (b * a) = 1 div b"
64240
eabf80376aab more standardized names
haftmann
parents: 64239
diff changeset
  1204
  using assms is_unit_div_mult_cancel_left [of a b] by (simp add: ac_simps)
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1205
64591
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
  1206
lemma unit_div_eq_0_iff:
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
  1207
  assumes "is_unit b"
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
  1208
  shows "a div b = 0 \<longleftrightarrow> a = 0"
75669
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75543
diff changeset
  1209
  using assms by (simp add: dvd_div_eq_0_iff unit_imp_dvd)
64591
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
  1210
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
  1211
lemma div_mult_unit2:
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
  1212
  "is_unit c \<Longrightarrow> b dvd a \<Longrightarrow> a div (b * c) = a div b div c"
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
  1213
  by (rule dvd_div_mult2_eq) (simp_all add: mult_unit_dvd_iff)
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
  1214
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1215
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1216
text \<open>Coprimality\<close>
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1217
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1218
definition coprime :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1219
  where "coprime a b \<longleftrightarrow> (\<forall>c. c dvd a \<longrightarrow> c dvd b \<longrightarrow> is_unit c)"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1220
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1221
lemma coprimeI:
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1222
  assumes "\<And>c. c dvd a \<Longrightarrow> c dvd b \<Longrightarrow> is_unit c"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1223
  shows "coprime a b"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1224
  using assms by (auto simp: coprime_def)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1225
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1226
lemma not_coprimeI:
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1227
  assumes "c dvd a" and "c dvd b" and "\<not> is_unit c"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1228
  shows "\<not> coprime a b"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1229
  using assms by (auto simp: coprime_def)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1230
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1231
lemma coprime_common_divisor:
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1232
  "is_unit c" if "coprime a b" and "c dvd a" and "c dvd b"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1233
  using that by (auto simp: coprime_def)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1234
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1235
lemma not_coprimeE:
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1236
  assumes "\<not> coprime a b"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1237
  obtains c where "c dvd a" and "c dvd b" and "\<not> is_unit c"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1238
  using assms by (auto simp: coprime_def)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1239
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1240
lemma coprime_imp_coprime:
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1241
  "coprime a b" if "coprime c d"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1242
    and "\<And>e. \<not> is_unit e \<Longrightarrow> e dvd a \<Longrightarrow> e dvd b \<Longrightarrow> e dvd c"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1243
    and "\<And>e. \<not> is_unit e \<Longrightarrow> e dvd a \<Longrightarrow> e dvd b \<Longrightarrow> e dvd d"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1244
proof (rule coprimeI)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1245
  fix e
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1246
  assume "e dvd a" and "e dvd b"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1247
  with that have "e dvd c" and "e dvd d"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1248
    by (auto intro: dvd_trans)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1249
  with \<open>coprime c d\<close> show "is_unit e"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1250
    by (rule coprime_common_divisor)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1251
qed
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1252
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1253
lemma coprime_divisors:
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1254
  "coprime a b" if "a dvd c" "b dvd d" and "coprime c d"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1255
using \<open>coprime c d\<close> proof (rule coprime_imp_coprime)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1256
  fix e
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1257
  assume "e dvd a" then show "e dvd c"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1258
    using \<open>a dvd c\<close> by (rule dvd_trans)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1259
  assume "e dvd b" then show "e dvd d"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1260
    using \<open>b dvd d\<close> by (rule dvd_trans)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1261
qed
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1262
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1263
lemma coprime_self [simp]:
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1264
  "coprime a a \<longleftrightarrow> is_unit a" (is "?P \<longleftrightarrow> ?Q")
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1265
proof
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1266
  assume ?P
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1267
  then show ?Q
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1268
    by (rule coprime_common_divisor) simp_all
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1269
next
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1270
  assume ?Q
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1271
  show ?P
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1272
    by (rule coprimeI) (erule dvd_unit_imp_unit, rule \<open>?Q\<close>)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1273
qed
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1274
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1275
lemma coprime_commute [ac_simps]:
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1276
  "coprime b a \<longleftrightarrow> coprime a b"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1277
  unfolding coprime_def by auto
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1278
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1279
lemma is_unit_left_imp_coprime:
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1280
  "coprime a b" if "is_unit a"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1281
proof (rule coprimeI)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1282
  fix c
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1283
  assume "c dvd a"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1284
  with that show "is_unit c"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1285
    by (auto intro: dvd_unit_imp_unit)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1286
qed
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1287
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1288
lemma is_unit_right_imp_coprime:
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1289
  "coprime a b" if "is_unit b"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1290
  using that is_unit_left_imp_coprime [of b a] by (simp add: ac_simps)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1291
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1292
lemma coprime_1_left [simp]:
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1293
  "coprime 1 a"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1294
  by (rule coprimeI)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1295
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1296
lemma coprime_1_right [simp]:
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1297
  "coprime a 1"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1298
  by (rule coprimeI)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1299
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1300
lemma coprime_0_left_iff [simp]:
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1301
  "coprime 0 a \<longleftrightarrow> is_unit a"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1302
  by (auto intro: coprimeI dvd_unit_imp_unit coprime_common_divisor [of 0 a a])
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1303
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1304
lemma coprime_0_right_iff [simp]:
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1305
  "coprime a 0 \<longleftrightarrow> is_unit a"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1306
  using coprime_0_left_iff [of a] by (simp add: ac_simps)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1307
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1308
lemma coprime_mult_self_left_iff [simp]:
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1309
  "coprime (c * a) (c * b) \<longleftrightarrow> is_unit c \<and> coprime a b"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1310
  by (auto intro: coprime_common_divisor)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1311
    (rule coprimeI, auto intro: coprime_common_divisor simp add: dvd_mult_unit_iff')+
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1312
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1313
lemma coprime_mult_self_right_iff [simp]:
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1314
  "coprime (a * c) (b * c) \<longleftrightarrow> is_unit c \<and> coprime a b"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1315
  using coprime_mult_self_left_iff [of c a b] by (simp add: ac_simps)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1316
67234
ab10ea1d6fd0 Some lemmas on complex numbers and coprimality
eberlm <eberlm@in.tum.de>
parents: 67226
diff changeset
  1317
lemma coprime_absorb_left:
ab10ea1d6fd0 Some lemmas on complex numbers and coprimality
eberlm <eberlm@in.tum.de>
parents: 67226
diff changeset
  1318
  assumes "x dvd y"
ab10ea1d6fd0 Some lemmas on complex numbers and coprimality
eberlm <eberlm@in.tum.de>
parents: 67226
diff changeset
  1319
  shows   "coprime x y \<longleftrightarrow> is_unit x"
ab10ea1d6fd0 Some lemmas on complex numbers and coprimality
eberlm <eberlm@in.tum.de>
parents: 67226
diff changeset
  1320
  using assms coprime_common_divisor is_unit_left_imp_coprime by auto
ab10ea1d6fd0 Some lemmas on complex numbers and coprimality
eberlm <eberlm@in.tum.de>
parents: 67226
diff changeset
  1321
ab10ea1d6fd0 Some lemmas on complex numbers and coprimality
eberlm <eberlm@in.tum.de>
parents: 67226
diff changeset
  1322
lemma coprime_absorb_right:
ab10ea1d6fd0 Some lemmas on complex numbers and coprimality
eberlm <eberlm@in.tum.de>
parents: 67226
diff changeset
  1323
  assumes "y dvd x"
ab10ea1d6fd0 Some lemmas on complex numbers and coprimality
eberlm <eberlm@in.tum.de>
parents: 67226
diff changeset
  1324
  shows   "coprime x y \<longleftrightarrow> is_unit y"
ab10ea1d6fd0 Some lemmas on complex numbers and coprimality
eberlm <eberlm@in.tum.de>
parents: 67226
diff changeset
  1325
  using assms coprime_common_divisor is_unit_right_imp_coprime by auto
ab10ea1d6fd0 Some lemmas on complex numbers and coprimality
eberlm <eberlm@in.tum.de>
parents: 67226
diff changeset
  1326
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1327
end
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1328
64848
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1329
class unit_factor =
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1330
  fixes unit_factor :: "'a \<Rightarrow> 'a"
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1331
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1332
class semidom_divide_unit_factor = semidom_divide + unit_factor +
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1333
  assumes unit_factor_0 [simp]: "unit_factor 0 = 0"
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1334
    and is_unit_unit_factor: "a dvd 1 \<Longrightarrow> unit_factor a = a"
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1335
    and unit_factor_is_unit: "a \<noteq> 0 \<Longrightarrow> unit_factor a dvd 1"
71398
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1336
    and unit_factor_mult_unit_left: "a dvd 1 \<Longrightarrow> unit_factor (a * b) = a * unit_factor b"
67226
ec32cdaab97b isabelle update_cartouches -c -t;
wenzelm
parents: 67084
diff changeset
  1337
  \<comment> \<open>This fine-grained hierarchy will later on allow lean normalization of polynomials\<close>
71398
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1338
begin
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1339
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1340
lemma unit_factor_mult_unit_right: "a dvd 1 \<Longrightarrow> unit_factor (b * a) = unit_factor b * a"
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1341
  using unit_factor_mult_unit_left[of a b] by (simp add: mult_ac)
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1342
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1343
lemmas [simp] = unit_factor_mult_unit_left unit_factor_mult_unit_right
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1344
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1345
end
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1346
64848
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1347
class normalization_semidom = algebraic_semidom + semidom_divide_unit_factor +
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1348
  fixes normalize :: "'a \<Rightarrow> 'a"
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1349
  assumes unit_factor_mult_normalize [simp]: "unit_factor a * normalize a = a"
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  1350
    and normalize_0 [simp]: "normalize 0 = 0"
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1351
begin
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1352
60688
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1353
text \<open>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68253
diff changeset
  1354
  Class \<^class>\<open>normalization_semidom\<close> cultivates the idea that each integral
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  1355
  domain can be split into equivalence classes whose representants are
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68253
diff changeset
  1356
  associated, i.e. divide each other. \<^const>\<open>normalize\<close> specifies a canonical
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  1357
  representant for each equivalence class. The rationale behind this is that
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  1358
  it is easier to reason about equality than equivalences, hence we prefer to
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  1359
  think about equality of normalized values rather than associated elements.
60688
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1360
\<close>
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1361
64848
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1362
declare unit_factor_is_unit [iff]
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1363
  
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1364
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
  1365
  by (rule unit_imp_dvd) simp
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1366
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1367
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
  1368
  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
  1369
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1370
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
  1371
  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
  1372
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1373
lemma normalize_eq_0_iff [simp]: "normalize a = 0 \<longleftrightarrow> a = 0"
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  1374
  (is "?lhs \<longleftrightarrow> ?rhs")
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1375
proof
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  1376
  assume ?lhs
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1377
  moreover have "unit_factor a * normalize a = a" by simp
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  1378
  ultimately show ?rhs by simp
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1379
next
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  1380
  assume ?rhs
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  1381
  then show ?lhs by simp
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1382
qed
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1383
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1384
lemma unit_factor_eq_0_iff [simp]: "unit_factor a = 0 \<longleftrightarrow> a = 0"
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  1385
  (is "?lhs \<longleftrightarrow> ?rhs")
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1386
proof
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  1387
  assume ?lhs
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1388
  moreover have "unit_factor a * normalize a = a" by simp
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  1389
  ultimately show ?rhs by simp
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1390
next
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  1391
  assume ?rhs
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  1392
  then show ?lhs by simp
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1393
qed
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1394
64848
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1395
lemma div_unit_factor [simp]: "a div unit_factor a = normalize a"
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1396
proof (cases "a = 0")
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1397
  case True
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1398
  then show ?thesis by simp
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1399
next
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1400
  case False
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1401
  then have "unit_factor a \<noteq> 0"
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1402
    by simp
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1403
  with nonzero_mult_div_cancel_left
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1404
  have "unit_factor a * normalize a div unit_factor a = normalize a"
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1405
    by blast
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1406
  then show ?thesis by simp
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1407
qed
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1408
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1409
lemma normalize_div [simp]: "normalize a div a = 1 div unit_factor a"
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1410
proof (cases "a = 0")
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1411
  case True
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1412
  then show ?thesis by simp
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1413
next
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1414
  case False
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1415
  have "normalize a div a = normalize a div (unit_factor a * normalize a)"
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1416
    by simp
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1417
  also have "\<dots> = 1 div unit_factor a"
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1418
    using False by (subst is_unit_div_mult_cancel_right) simp_all
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1419
  finally show ?thesis .
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1420
qed
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1421
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1422
lemma is_unit_normalize:
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1423
  assumes "is_unit a"
64848
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1424
  shows "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
  1425
proof -
64848
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1426
  from assms have "unit_factor a = a"
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1427
    by (rule is_unit_unit_factor)
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1428
  moreover from assms have "a \<noteq> 0"
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1429
    by auto
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1430
  moreover have "normalize a = a div unit_factor a"
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1431
    by simp
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1432
  ultimately show ?thesis
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  1433
    by simp
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1434
qed
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1435
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1436
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
  1437
  by (rule is_unit_unit_factor) simp
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1438
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1439
lemma normalize_1 [simp]: "normalize 1 = 1"
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1440
  by (rule is_unit_normalize) simp
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1441
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1442
lemma normalize_1_iff: "normalize a = 1 \<longleftrightarrow> is_unit a"
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  1443
  (is "?lhs \<longleftrightarrow> ?rhs")
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1444
proof
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  1445
  assume ?rhs
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  1446
  then show ?lhs by (rule is_unit_normalize)
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1447
next
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  1448
  assume ?lhs
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  1449
  then have "unit_factor a * normalize a = unit_factor a * 1"
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1450
    by simp
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1451
  then have "unit_factor a = a"
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1452
    by simp
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  1453
  moreover
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  1454
  from \<open>?lhs\<close> have "a \<noteq> 0" by auto
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  1455
  then have "is_unit (unit_factor a)" by simp
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  1456
  ultimately show ?rhs by simp
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1457
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
  1458
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1459
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
  1460
proof (cases "a = 0")
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1461
  case True
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1462
  then show ?thesis by simp
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1463
next
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1464
  case False
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1465
  then have "normalize a \<noteq> 0" by simp
64240
eabf80376aab more standardized names
haftmann
parents: 64239
diff changeset
  1466
  with nonzero_mult_div_cancel_right
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1467
  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
  1468
  then show ?thesis by simp
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1469
qed
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1470
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1471
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
  1472
  by (cases "b = 0") simp_all
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1473
63947
559f0882d6a6 more lemmas
haftmann
parents: 63924
diff changeset
  1474
lemma inv_unit_factor_eq_0_iff [simp]:
559f0882d6a6 more lemmas
haftmann
parents: 63924
diff changeset
  1475
  "1 div unit_factor a = 0 \<longleftrightarrow> a = 0"
559f0882d6a6 more lemmas
haftmann
parents: 63924
diff changeset
  1476
  (is "?lhs \<longleftrightarrow> ?rhs")
559f0882d6a6 more lemmas
haftmann
parents: 63924
diff changeset
  1477
proof
559f0882d6a6 more lemmas
haftmann
parents: 63924
diff changeset
  1478
  assume ?lhs
559f0882d6a6 more lemmas
haftmann
parents: 63924
diff changeset
  1479
  then have "a * (1 div unit_factor a) = a * 0"
559f0882d6a6 more lemmas
haftmann
parents: 63924
diff changeset
  1480
    by simp
559f0882d6a6 more lemmas
haftmann
parents: 63924
diff changeset
  1481
  then show ?rhs
559f0882d6a6 more lemmas
haftmann
parents: 63924
diff changeset
  1482
    by simp
559f0882d6a6 more lemmas
haftmann
parents: 63924
diff changeset
  1483
next
559f0882d6a6 more lemmas
haftmann
parents: 63924
diff changeset
  1484
  assume ?rhs
559f0882d6a6 more lemmas
haftmann
parents: 63924
diff changeset
  1485
  then show ?lhs by simp
559f0882d6a6 more lemmas
haftmann
parents: 63924
diff changeset
  1486
qed
559f0882d6a6 more lemmas
haftmann
parents: 63924
diff changeset
  1487
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1488
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
  1489
  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
  1490
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1491
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
  1492
  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
  1493
71398
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1494
lemma normalize_mult_unit_left [simp]:
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1495
  assumes "a dvd 1"
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1496
  shows   "normalize (a * b) = normalize b"
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1497
proof (cases "b = 0")
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1498
  case False
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1499
  have "a * unit_factor b * normalize (a * b) = unit_factor (a * b) * normalize (a * b)"
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1500
    using assms by (subst unit_factor_mult_unit_left) auto
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1501
  also have "\<dots> = a * b" by simp
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1502
  also have "b = unit_factor b * normalize b" by simp
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1503
  hence "a * b = a * unit_factor b * normalize b"
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1504
    by (simp only: mult_ac)
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1505
  finally show ?thesis
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1506
    using assms False by auto
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1507
qed auto
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1508
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1509
lemma normalize_mult_unit_right [simp]:
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1510
  assumes "b dvd 1"
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1511
  shows   "normalize (a * b) = normalize a"
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1512
  using assms by (subst mult.commute) auto
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1513
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1514
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
  1515
proof (cases "a = 0")
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1516
  case False
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1517
  have "normalize a = normalize (unit_factor a * normalize a)"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1518
    by simp
71398
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1519
  also from False have "\<dots> = normalize (normalize a)"
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1520
    by (subst normalize_mult_unit_left) auto
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1521
  finally show ?thesis ..
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1522
qed auto
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1523
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1524
lemma unit_factor_normalize [simp]:
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1525
  assumes "a \<noteq> 0"
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1526
  shows "unit_factor (normalize a) = 1"
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1527
proof -
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1528
  from assms have *: "normalize a \<noteq> 0"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1529
    by simp
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1530
  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
  1531
    by (simp only: unit_factor_mult_normalize)
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1532
  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
  1533
    by simp
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1534
  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
  1535
    by simp
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1536
  with * show ?thesis
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1537
    by simp
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1538
qed
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1539
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1540
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
  1541
proof -
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1542
  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
  1543
    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
  1544
      by (cases "a = 0") simp_all
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1545
  then show ?thesis by simp
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1546
qed
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1547
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1548
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
  1549
proof -
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1550
  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
  1551
    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
  1552
      by (cases "b = 0") simp_all
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1553
  then show ?thesis by simp
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1554
qed
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1555
65811
2653f1cd8775 more lemmas
haftmann
parents: 64848
diff changeset
  1556
lemma normalize_idem_imp_unit_factor_eq:
2653f1cd8775 more lemmas
haftmann
parents: 64848
diff changeset
  1557
  assumes "normalize a = a"
2653f1cd8775 more lemmas
haftmann
parents: 64848
diff changeset
  1558
  shows "unit_factor a = of_bool (a \<noteq> 0)"
2653f1cd8775 more lemmas
haftmann
parents: 64848
diff changeset
  1559
proof (cases "a = 0")
2653f1cd8775 more lemmas
haftmann
parents: 64848
diff changeset
  1560
  case True
2653f1cd8775 more lemmas
haftmann
parents: 64848
diff changeset
  1561
  then show ?thesis
2653f1cd8775 more lemmas
haftmann
parents: 64848
diff changeset
  1562
    by simp
2653f1cd8775 more lemmas
haftmann
parents: 64848
diff changeset
  1563
next
2653f1cd8775 more lemmas
haftmann
parents: 64848
diff changeset
  1564
  case False
2653f1cd8775 more lemmas
haftmann
parents: 64848
diff changeset
  1565
  then show ?thesis
2653f1cd8775 more lemmas
haftmann
parents: 64848
diff changeset
  1566
    using assms unit_factor_normalize [of a] by simp
2653f1cd8775 more lemmas
haftmann
parents: 64848
diff changeset
  1567
qed
2653f1cd8775 more lemmas
haftmann
parents: 64848
diff changeset
  1568
2653f1cd8775 more lemmas
haftmann
parents: 64848
diff changeset
  1569
lemma normalize_idem_imp_is_unit_iff:
2653f1cd8775 more lemmas
haftmann
parents: 64848
diff changeset
  1570
  assumes "normalize a = a"
2653f1cd8775 more lemmas
haftmann
parents: 64848
diff changeset
  1571
  shows "is_unit a \<longleftrightarrow> a = 1"
2653f1cd8775 more lemmas
haftmann
parents: 64848
diff changeset
  1572
  using assms by (cases "a = 0") (auto dest: is_unit_normalize)
2653f1cd8775 more lemmas
haftmann
parents: 64848
diff changeset
  1573
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1574
lemma coprime_normalize_left_iff [simp]:
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1575
  "coprime (normalize a) b \<longleftrightarrow> coprime a b"
75669
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75543
diff changeset
  1576
  by (rule iffI; rule coprimeI) (auto intro: coprime_common_divisor)
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1577
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1578
lemma coprime_normalize_right_iff [simp]:
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1579
  "coprime a (normalize b) \<longleftrightarrow> coprime a b"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1580
  using coprime_normalize_left_iff [of b a] by (simp add: ac_simps)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  1581
60688
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1582
text \<open>
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  1583
  We avoid an explicit definition of associated elements but prefer explicit
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68253
diff changeset
  1584
  normalisation instead. In theory we could define an abbreviation like \<^prop>\<open>associated a b \<longleftrightarrow> normalize a = normalize b\<close> but this is counterproductive
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  1585
  without suggestive infix syntax, which we do not want to sacrifice for this
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  1586
  purpose here.
60688
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1587
\<close>
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1588
60688
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1589
lemma associatedI:
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1590
  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
  1591
  shows "normalize a = normalize b"
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1592
proof (cases "a = 0 \<or> b = 0")
63325
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 assms show ?thesis by auto
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1595
next
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1596
  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
  1597
  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
  1598
  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
  1599
  ultimately have "b * 1 = b * (c * d)"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1600
    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
  1601
  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
  1602
    unfolding mult_cancel_left by simp
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1603
  then have "is_unit c" and "is_unit d"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1604
    by auto
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1605
  with a b show ?thesis
71398
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1606
    by (simp add: 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
  1607
qed
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1608
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1609
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
  1610
  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
  1611
  by simp
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1612
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1613
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
  1614
  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
  1615
  by simp
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1616
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1617
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
  1618
  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
  1619
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1620
lemma associated_iff_dvd: "normalize a = normalize b \<longleftrightarrow> a dvd b \<and> b dvd a"
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  1621
  (is "?lhs \<longleftrightarrow> ?rhs")
60688
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1622
proof
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  1623
  assume ?rhs
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  1624
  then show ?lhs 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
  1625
next
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  1626
  assume ?lhs
60688
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1627
  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
  1628
    by simp
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1629
  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
  1630
    by (simp add: ac_simps)
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  1631
  show ?rhs
60688
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1632
  proof (cases "a = 0 \<or> b = 0")
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1633
    case True
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  1634
    with \<open>?lhs\<close> show ?thesis by auto
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1635
  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
  1636
    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
  1637
    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
  1638
      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
  1639
    with * show ?thesis by simp
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1640
  qed
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1641
qed
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1642
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1643
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
  1644
  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
  1645
  assumes "normalize a = a" and "normalize b = b"
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1646
  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
  1647
proof -
01488b559910 avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
haftmann
parents: 60685
diff changeset
  1648
  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
  1649
    unfolding associated_iff_dvd by simp
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  1650
  with \<open>normalize a = a\<close> have "a = normalize b"
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  1651
    by simp
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  1652
  with \<open>normalize b = b\<close> show "a = b"
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  1653
    by simp
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1654
qed
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1655
64591
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
  1656
lemma normalize_unit_factor_eqI:
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
  1657
  assumes "normalize a = normalize b"
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
  1658
    and "unit_factor a = unit_factor b"
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
  1659
  shows "a = b"
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
  1660
proof -
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
  1661
  from assms have "unit_factor a * normalize a = unit_factor b * normalize b"
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
  1662
    by simp
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
  1663
  then show ?thesis
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
  1664
    by simp
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
  1665
qed
240a39af9ec4 restructured matter on polynomials and normalized fractions
haftmann
parents: 64290
diff changeset
  1666
71398
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1667
lemma normalize_mult_normalize_left [simp]: "normalize (normalize a * b) = normalize (a * b)"
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1668
  by (rule associated_eqI) (auto intro!: mult_dvd_mono)
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1669
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1670
lemma normalize_mult_normalize_right [simp]: "normalize (a * normalize b) = normalize (a * b)"
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1671
  by (rule associated_eqI) (auto intro!: mult_dvd_mono)
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1672
60685
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1673
end
cb21b7022b00 moved normalization and unit_factor into Main HOL corpus
haftmann
parents: 60615
diff changeset
  1674
64164
38c407446400 separate type class for arbitrary quotient and remainder partitions
haftmann
parents: 63950
diff changeset
  1675
71398
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1676
class normalization_semidom_multiplicative = normalization_semidom +
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1677
  assumes unit_factor_mult: "unit_factor (a * b) = unit_factor a * unit_factor b"
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1678
begin
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1679
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1680
lemma normalize_mult: "normalize (a * b) = normalize a * normalize b"
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1681
proof (cases "a = 0 \<or> b = 0")
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1682
  case True
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1683
  then show ?thesis by auto
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1684
next
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1685
  case False
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1686
  have "unit_factor (a * b) * normalize (a * b) = a * b"
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1687
    by (rule unit_factor_mult_normalize)
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1688
  then have "normalize (a * b) = a * b div unit_factor (a * b)"
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1689
    by simp
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1690
  also have "\<dots> = a * b div unit_factor (b * a)"
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1691
    by (simp add: ac_simps)
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1692
  also have "\<dots> = a * b div unit_factor b div unit_factor a"
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1693
    using False by (simp add: unit_factor_mult is_unit_div_mult2_eq [symmetric])
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1694
  also have "\<dots> = a * (b div unit_factor b) div unit_factor a"
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1695
    using False by (subst unit_div_mult_swap) simp_all
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1696
  also have "\<dots> = normalize a * normalize b"
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1697
    using False
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1698
    by (simp add: mult.commute [of a] mult.commute [of "normalize a"] unit_div_mult_swap [symmetric])
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1699
  finally show ?thesis .
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1700
qed
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1701
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1702
lemma dvd_unit_factor_div:
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1703
  assumes "b dvd a"
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1704
  shows "unit_factor (a div b) = unit_factor a div unit_factor b"
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1705
proof -
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1706
  from assms have "a = a div b * b"
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1707
    by simp
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1708
  then have "unit_factor a = unit_factor (a div b * b)"
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1709
    by simp
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1710
  then show ?thesis
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1711
    by (cases "b = 0") (simp_all add: unit_factor_mult)
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1712
qed
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1713
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1714
lemma dvd_normalize_div:
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1715
  assumes "b dvd a"
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1716
  shows "normalize (a div b) = normalize a div normalize b"
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1717
proof -
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1718
  from assms have "a = a div b * b"
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1719
    by simp
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1720
  then have "normalize a = normalize (a div b * b)"
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1721
    by simp
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1722
  then show ?thesis
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1723
    by (cases "b = 0") (simp_all add: normalize_mult)
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1724
qed
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1725
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1726
end
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1727
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1728
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1729
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
  1730
64164
38c407446400 separate type class for arbitrary quotient and remainder partitions
haftmann
parents: 63950
diff changeset
  1731
text \<open>Syntactic division remainder operator\<close>
38c407446400 separate type class for arbitrary quotient and remainder partitions
haftmann
parents: 63950
diff changeset
  1732
38c407446400 separate type class for arbitrary quotient and remainder partitions
haftmann
parents: 63950
diff changeset
  1733
class modulo = dvd + divide +
80932
261cd8722677 standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
wenzelm
parents: 79566
diff changeset
  1734
  fixes modulo :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  (infixl \<open>mod\<close> 70)
64164
38c407446400 separate type class for arbitrary quotient and remainder partitions
haftmann
parents: 63950
diff changeset
  1735
38c407446400 separate type class for arbitrary quotient and remainder partitions
haftmann
parents: 63950
diff changeset
  1736
text \<open>Arbitrary quotient and remainder partitions\<close>
38c407446400 separate type class for arbitrary quotient and remainder partitions
haftmann
parents: 63950
diff changeset
  1737
38c407446400 separate type class for arbitrary quotient and remainder partitions
haftmann
parents: 63950
diff changeset
  1738
class semiring_modulo = comm_semiring_1_cancel + divide + modulo +
79541
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
  1739
  assumes div_mult_mod_eq: \<open>a div b * b + a mod b = a\<close>
64164
38c407446400 separate type class for arbitrary quotient and remainder partitions
haftmann
parents: 63950
diff changeset
  1740
begin
38c407446400 separate type class for arbitrary quotient and remainder partitions
haftmann
parents: 63950
diff changeset
  1741
38c407446400 separate type class for arbitrary quotient and remainder partitions
haftmann
parents: 63950
diff changeset
  1742
lemma mod_div_decomp:
38c407446400 separate type class for arbitrary quotient and remainder partitions
haftmann
parents: 63950
diff changeset
  1743
  fixes a b
38c407446400 separate type class for arbitrary quotient and remainder partitions
haftmann
parents: 63950
diff changeset
  1744
  obtains q r where "q = a div b" and "r = a mod b"
38c407446400 separate type class for arbitrary quotient and remainder partitions
haftmann
parents: 63950
diff changeset
  1745
    and "a = q * b + r"
38c407446400 separate type class for arbitrary quotient and remainder partitions
haftmann
parents: 63950
diff changeset
  1746
proof -
64242
93c6f0da5c70 more standardized theorem names for facts involving the div and mod identity
haftmann
parents: 64240
diff changeset
  1747
  from div_mult_mod_eq have "a = a div b * b + a mod b" by simp
64164
38c407446400 separate type class for arbitrary quotient and remainder partitions
haftmann
parents: 63950
diff changeset
  1748
  moreover have "a div b = a div b" ..
38c407446400 separate type class for arbitrary quotient and remainder partitions
haftmann
parents: 63950
diff changeset
  1749
  moreover have "a mod b = a mod b" ..
38c407446400 separate type class for arbitrary quotient and remainder partitions
haftmann
parents: 63950
diff changeset
  1750
  note that ultimately show thesis by blast
38c407446400 separate type class for arbitrary quotient and remainder partitions
haftmann
parents: 63950
diff changeset
  1751
qed
38c407446400 separate type class for arbitrary quotient and remainder partitions
haftmann
parents: 63950
diff changeset
  1752
64242
93c6f0da5c70 more standardized theorem names for facts involving the div and mod identity
haftmann
parents: 64240
diff changeset
  1753
lemma mult_div_mod_eq: "b * (a div b) + a mod b = a"
93c6f0da5c70 more standardized theorem names for facts involving the div and mod identity
haftmann
parents: 64240
diff changeset
  1754
  using div_mult_mod_eq [of a b] by (simp add: ac_simps)
64164
38c407446400 separate type class for arbitrary quotient and remainder partitions
haftmann
parents: 63950
diff changeset
  1755
64242
93c6f0da5c70 more standardized theorem names for facts involving the div and mod identity
haftmann
parents: 64240
diff changeset
  1756
lemma mod_div_mult_eq: "a mod b + a div b * b = a"
93c6f0da5c70 more standardized theorem names for facts involving the div and mod identity
haftmann
parents: 64240
diff changeset
  1757
  using div_mult_mod_eq [of a b] by (simp add: ac_simps)
64164
38c407446400 separate type class for arbitrary quotient and remainder partitions
haftmann
parents: 63950
diff changeset
  1758
64242
93c6f0da5c70 more standardized theorem names for facts involving the div and mod identity
haftmann
parents: 64240
diff changeset
  1759
lemma mod_mult_div_eq: "a mod b + b * (a div b) = a"
93c6f0da5c70 more standardized theorem names for facts involving the div and mod identity
haftmann
parents: 64240
diff changeset
  1760
  using div_mult_mod_eq [of a b] by (simp add: ac_simps)
64164
38c407446400 separate type class for arbitrary quotient and remainder partitions
haftmann
parents: 63950
diff changeset
  1761
64242
93c6f0da5c70 more standardized theorem names for facts involving the div and mod identity
haftmann
parents: 64240
diff changeset
  1762
lemma minus_div_mult_eq_mod: "a - a div b * b = a mod b"
93c6f0da5c70 more standardized theorem names for facts involving the div and mod identity
haftmann
parents: 64240
diff changeset
  1763
  by (rule add_implies_diff [symmetric]) (fact mod_div_mult_eq)
64164
38c407446400 separate type class for arbitrary quotient and remainder partitions
haftmann
parents: 63950
diff changeset
  1764
64242
93c6f0da5c70 more standardized theorem names for facts involving the div and mod identity
haftmann
parents: 64240
diff changeset
  1765
lemma minus_mult_div_eq_mod: "a - b * (a div b) = a mod b"
93c6f0da5c70 more standardized theorem names for facts involving the div and mod identity
haftmann
parents: 64240
diff changeset
  1766
  by (rule add_implies_diff [symmetric]) (fact mod_mult_div_eq)
64164
38c407446400 separate type class for arbitrary quotient and remainder partitions
haftmann
parents: 63950
diff changeset
  1767
64242
93c6f0da5c70 more standardized theorem names for facts involving the div and mod identity
haftmann
parents: 64240
diff changeset
  1768
lemma minus_mod_eq_div_mult: "a - a mod b = a div b * b"
93c6f0da5c70 more standardized theorem names for facts involving the div and mod identity
haftmann
parents: 64240
diff changeset
  1769
  by (rule add_implies_diff [symmetric]) (fact div_mult_mod_eq)
64164
38c407446400 separate type class for arbitrary quotient and remainder partitions
haftmann
parents: 63950
diff changeset
  1770
64242
93c6f0da5c70 more standardized theorem names for facts involving the div and mod identity
haftmann
parents: 64240
diff changeset
  1771
lemma minus_mod_eq_mult_div: "a - a mod b = b * (a div b)"
93c6f0da5c70 more standardized theorem names for facts involving the div and mod identity
haftmann
parents: 64240
diff changeset
  1772
  by (rule add_implies_diff [symmetric]) (fact mult_div_mod_eq)
64164
38c407446400 separate type class for arbitrary quotient and remainder partitions
haftmann
parents: 63950
diff changeset
  1773
70902
cb161182ce7f generalized
haftmann
parents: 70817
diff changeset
  1774
lemma mod_0_imp_dvd [dest!]:
cb161182ce7f generalized
haftmann
parents: 70817
diff changeset
  1775
  "b dvd a" if "a mod b = 0"
cb161182ce7f generalized
haftmann
parents: 70817
diff changeset
  1776
proof -
cb161182ce7f generalized
haftmann
parents: 70817
diff changeset
  1777
  have "b dvd (a div b) * b" by simp
cb161182ce7f generalized
haftmann
parents: 70817
diff changeset
  1778
  also have "(a div b) * b = a"
cb161182ce7f generalized
haftmann
parents: 70817
diff changeset
  1779
    using div_mult_mod_eq [of a b] by (simp add: that)
cb161182ce7f generalized
haftmann
parents: 70817
diff changeset
  1780
  finally show ?thesis .
cb161182ce7f generalized
haftmann
parents: 70817
diff changeset
  1781
qed
cb161182ce7f generalized
haftmann
parents: 70817
diff changeset
  1782
68253
a8660a39e304 grouped material on numeral division
haftmann
parents: 68252
diff changeset
  1783
lemma [nitpick_unfold]:
a8660a39e304 grouped material on numeral division
haftmann
parents: 68252
diff changeset
  1784
  "a mod b = a - a div b * b"
a8660a39e304 grouped material on numeral division
haftmann
parents: 68252
diff changeset
  1785
  by (fact minus_div_mult_eq_mod [symmetric])
a8660a39e304 grouped material on numeral division
haftmann
parents: 68252
diff changeset
  1786
64164
38c407446400 separate type class for arbitrary quotient and remainder partitions
haftmann
parents: 63950
diff changeset
  1787
end
64242
93c6f0da5c70 more standardized theorem names for facts involving the div and mod identity
haftmann
parents: 64240
diff changeset
  1788
79541
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
  1789
class semiring_modulo_trivial = semiring_modulo + divide_trivial
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
  1790
begin
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
  1791
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
  1792
lemma mod_0 [simp]:
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
  1793
  \<open>0 mod a = 0\<close>
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
  1794
  using div_mult_mod_eq [of 0 a] by simp
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
  1795
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
  1796
lemma mod_by_0 [simp]:
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
  1797
  \<open>a mod 0 = a\<close>
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
  1798
  using div_mult_mod_eq [of a 0] by simp
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
  1799
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
  1800
lemma mod_by_1 [simp]:
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
  1801
  \<open>a mod 1 = 0\<close>
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
  1802
proof -
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
  1803
  have \<open>a + a mod 1 = a\<close>
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
  1804
    using div_mult_mod_eq [of a 1] by simp
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
  1805
  then have \<open>a + a mod 1 = a + 0\<close>
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
  1806
    by simp
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
  1807
  then show ?thesis
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
  1808
    by (rule add_left_imp_eq)
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
  1809
qed
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
  1810
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
  1811
end
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
  1812
64164
38c407446400 separate type class for arbitrary quotient and remainder partitions
haftmann
parents: 63950
diff changeset
  1813
70145
f07b8fb99818 more document structure
haftmann
parents: 70144
diff changeset
  1814
subsection \<open>Quotient and remainder in integral domains\<close>
66807
c3631f32dfeb tuned structure
haftmann
parents: 66793
diff changeset
  1815
c3631f32dfeb tuned structure
haftmann
parents: 66793
diff changeset
  1816
class semidom_modulo = algebraic_semidom + semiring_modulo
c3631f32dfeb tuned structure
haftmann
parents: 66793
diff changeset
  1817
begin
c3631f32dfeb tuned structure
haftmann
parents: 66793
diff changeset
  1818
79541
4f40225936d1 common type class for trivial properties on div/mod
haftmann
parents: 78935
diff changeset
  1819
subclass semiring_modulo_trivial ..
66807
c3631f32dfeb tuned structure
haftmann
parents: 66793
diff changeset
  1820
c3631f32dfeb tuned structure
haftmann
parents: 66793
diff changeset
  1821
lemma mod_self [simp]:
c3631f32dfeb tuned structure
haftmann
parents: 66793
diff changeset
  1822
  "a mod a = 0"
c3631f32dfeb tuned structure
haftmann
parents: 66793
diff changeset
  1823
  using div_mult_mod_eq [of a a] by simp
c3631f32dfeb tuned structure
haftmann
parents: 66793
diff changeset
  1824
c3631f32dfeb tuned structure
haftmann
parents: 66793
diff changeset
  1825
lemma dvd_imp_mod_0 [simp]:
67084
haftmann
parents: 67051
diff changeset
  1826
  "b mod a = 0" if "a dvd b"
haftmann
parents: 67051
diff changeset
  1827
  using that minus_div_mult_eq_mod [of b a] by simp
66807
c3631f32dfeb tuned structure
haftmann
parents: 66793
diff changeset
  1828
c3631f32dfeb tuned structure
haftmann
parents: 66793
diff changeset
  1829
lemma mod_eq_0_iff_dvd:
c3631f32dfeb tuned structure
haftmann
parents: 66793
diff changeset
  1830
  "a mod b = 0 \<longleftrightarrow> b dvd a"
c3631f32dfeb tuned structure
haftmann
parents: 66793
diff changeset
  1831
  by (auto intro: mod_0_imp_dvd)
c3631f32dfeb tuned structure
haftmann
parents: 66793
diff changeset
  1832
c3631f32dfeb tuned structure
haftmann
parents: 66793
diff changeset
  1833
lemma dvd_eq_mod_eq_0 [nitpick_unfold, code]:
c3631f32dfeb tuned structure
haftmann
parents: 66793
diff changeset
  1834
  "a dvd b \<longleftrightarrow> b mod a = 0"
c3631f32dfeb tuned structure
haftmann
parents: 66793
diff changeset
  1835
  by (simp add: mod_eq_0_iff_dvd)
c3631f32dfeb tuned structure
haftmann
parents: 66793
diff changeset
  1836
c3631f32dfeb tuned structure
haftmann
parents: 66793
diff changeset
  1837
lemma dvd_mod_iff: 
c3631f32dfeb tuned structure
haftmann
parents: 66793
diff changeset
  1838
  assumes "c dvd b"
c3631f32dfeb tuned structure
haftmann
parents: 66793
diff changeset
  1839
  shows "c dvd a mod b \<longleftrightarrow> c dvd a"
c3631f32dfeb tuned structure
haftmann
parents: 66793
diff changeset
  1840
proof -
c3631f32dfeb tuned structure
haftmann
parents: 66793
diff changeset
  1841
  from assms have "(c dvd a mod b) \<longleftrightarrow> (c dvd ((a div b) * b + a mod b))" 
c3631f32dfeb tuned structure
haftmann
parents: 66793
diff changeset
  1842
    by (simp add: dvd_add_right_iff)
c3631f32dfeb tuned structure
haftmann
parents: 66793
diff changeset
  1843
  also have "(a div b) * b + a mod b = a"
c3631f32dfeb tuned structure
haftmann
parents: 66793
diff changeset
  1844
    using div_mult_mod_eq [of a b] by simp
c3631f32dfeb tuned structure
haftmann
parents: 66793
diff changeset
  1845
  finally show ?thesis .
c3631f32dfeb tuned structure
haftmann
parents: 66793
diff changeset
  1846
qed
c3631f32dfeb tuned structure
haftmann
parents: 66793
diff changeset
  1847
c3631f32dfeb tuned structure
haftmann
parents: 66793
diff changeset
  1848
lemma dvd_mod_imp_dvd:
c3631f32dfeb tuned structure
haftmann
parents: 66793
diff changeset
  1849
  assumes "c dvd a mod b" and "c dvd b"
c3631f32dfeb tuned structure
haftmann
parents: 66793
diff changeset
  1850
  shows "c dvd a"
c3631f32dfeb tuned structure
haftmann
parents: 66793
diff changeset
  1851
  using assms dvd_mod_iff [of c b a] by simp
c3631f32dfeb tuned structure
haftmann
parents: 66793
diff changeset
  1852
66808
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1853
lemma dvd_minus_mod [simp]:
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1854
  "b dvd a - a mod b"
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1855
  by (simp add: minus_mod_eq_div_mult)
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66807
diff changeset
  1856
66810
cc2b490f9dc4 generalized simproc
haftmann
parents: 66808
diff changeset
  1857
lemma cancel_div_mod_rules:
cc2b490f9dc4 generalized simproc
haftmann
parents: 66808
diff changeset
  1858
  "((a div b) * b + a mod b) + c = a + c"
cc2b490f9dc4 generalized simproc
haftmann
parents: 66808
diff changeset
  1859
  "(b * (a div b) + a mod b) + c = a + c"
cc2b490f9dc4 generalized simproc
haftmann
parents: 66808
diff changeset
  1860
  by (simp_all add: div_mult_mod_eq mult_div_mod_eq)
cc2b490f9dc4 generalized simproc
haftmann
parents: 66808
diff changeset
  1861
66807
c3631f32dfeb tuned structure
haftmann
parents: 66793
diff changeset
  1862
end
c3631f32dfeb tuned structure
haftmann
parents: 66793
diff changeset
  1863
70145
f07b8fb99818 more document structure
haftmann
parents: 70144
diff changeset
  1864
class idom_modulo = idom + semidom_modulo
f07b8fb99818 more document structure
haftmann
parents: 70144
diff changeset
  1865
begin
f07b8fb99818 more document structure
haftmann
parents: 70144
diff changeset
  1866
f07b8fb99818 more document structure
haftmann
parents: 70144
diff changeset
  1867
subclass idom_divide ..
f07b8fb99818 more document structure
haftmann
parents: 70144
diff changeset
  1868
f07b8fb99818 more document structure
haftmann
parents: 70144
diff changeset
  1869
lemma div_diff [simp]:
f07b8fb99818 more document structure
haftmann
parents: 70144
diff changeset
  1870
  "c dvd a \<Longrightarrow> c dvd b \<Longrightarrow> (a - b) div c = a div c - b div c"
f07b8fb99818 more document structure
haftmann
parents: 70144
diff changeset
  1871
  using div_add [of _  _ "- b"] by (simp add: dvd_neg_div)
f07b8fb99818 more document structure
haftmann
parents: 70144
diff changeset
  1872
f07b8fb99818 more document structure
haftmann
parents: 70144
diff changeset
  1873
end
f07b8fb99818 more document structure
haftmann
parents: 70144
diff changeset
  1874
f07b8fb99818 more document structure
haftmann
parents: 70144
diff changeset
  1875
f07b8fb99818 more document structure
haftmann
parents: 70144
diff changeset
  1876
subsection \<open>Interlude: basic tool support for algebraic and arithmetic calculations\<close>
66810
cc2b490f9dc4 generalized simproc
haftmann
parents: 66808
diff changeset
  1877
cc2b490f9dc4 generalized simproc
haftmann
parents: 66808
diff changeset
  1878
named_theorems arith "arith facts -- only ground formulas"
69605
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 69593
diff changeset
  1879
ML_file \<open>Tools/arith_data.ML\<close>
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 69593
diff changeset
  1880
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 69593
diff changeset
  1881
ML_file \<open>~~/src/Provers/Arith/cancel_div_mod.ML\<close>
66810
cc2b490f9dc4 generalized simproc
haftmann
parents: 66808
diff changeset
  1882
cc2b490f9dc4 generalized simproc
haftmann
parents: 66808
diff changeset
  1883
ML \<open>
cc2b490f9dc4 generalized simproc
haftmann
parents: 66808
diff changeset
  1884
structure Cancel_Div_Mod_Ring = Cancel_Div_Mod
cc2b490f9dc4 generalized simproc
haftmann
parents: 66808
diff changeset
  1885
(
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68253
diff changeset
  1886
  val div_name = \<^const_name>\<open>divide\<close>;
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68253
diff changeset
  1887
  val mod_name = \<^const_name>\<open>modulo\<close>;
66810
cc2b490f9dc4 generalized simproc
haftmann
parents: 66808
diff changeset
  1888
  val mk_binop = HOLogic.mk_binop;
cc2b490f9dc4 generalized simproc
haftmann
parents: 66808
diff changeset
  1889
  val mk_sum = Arith_Data.mk_sum;
cc2b490f9dc4 generalized simproc
haftmann
parents: 66808
diff changeset
  1890
  val dest_sum = Arith_Data.dest_sum;
cc2b490f9dc4 generalized simproc
haftmann
parents: 66808
diff changeset
  1891
cc2b490f9dc4 generalized simproc
haftmann
parents: 66808
diff changeset
  1892
  val div_mod_eqs = map mk_meta_eq @{thms cancel_div_mod_rules};
cc2b490f9dc4 generalized simproc
haftmann
parents: 66808
diff changeset
  1893
cc2b490f9dc4 generalized simproc
haftmann
parents: 66808
diff changeset
  1894
  val prove_eq_sums = Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac
cc2b490f9dc4 generalized simproc
haftmann
parents: 66808
diff changeset
  1895
    @{thms diff_conv_add_uminus add_0_left add_0_right ac_simps})
cc2b490f9dc4 generalized simproc
haftmann
parents: 66808
diff changeset
  1896
)
cc2b490f9dc4 generalized simproc
haftmann
parents: 66808
diff changeset
  1897
\<close>
cc2b490f9dc4 generalized simproc
haftmann
parents: 66808
diff changeset
  1898
cc2b490f9dc4 generalized simproc
haftmann
parents: 66808
diff changeset
  1899
simproc_setup cancel_div_mod_int ("(a::'a::semidom_modulo) + b") =
cc2b490f9dc4 generalized simproc
haftmann
parents: 66808
diff changeset
  1900
  \<open>K Cancel_Div_Mod_Ring.proc\<close>
cc2b490f9dc4 generalized simproc
haftmann
parents: 66808
diff changeset
  1901
70145
f07b8fb99818 more document structure
haftmann
parents: 70144
diff changeset
  1902
f07b8fb99818 more document structure
haftmann
parents: 70144
diff changeset
  1903
subsection \<open>Ordered semirings and rings\<close>
f07b8fb99818 more document structure
haftmann
parents: 70144
diff changeset
  1904
f07b8fb99818 more document structure
haftmann
parents: 70144
diff changeset
  1905
text \<open>
f07b8fb99818 more document structure
haftmann
parents: 70144
diff changeset
  1906
  The theory of partially ordered rings is taken from the books:
f07b8fb99818 more document structure
haftmann
parents: 70144
diff changeset
  1907
    \<^item> \<^emph>\<open>Lattice Theory\<close> by Garret Birkhoff, American Mathematical Society, 1979
f07b8fb99818 more document structure
haftmann
parents: 70144
diff changeset
  1908
    \<^item> \<^emph>\<open>Partially Ordered Algebraic Systems\<close>, Pergamon Press, 1963
f07b8fb99818 more document structure
haftmann
parents: 70144
diff changeset
  1909
f07b8fb99818 more document structure
haftmann
parents: 70144
diff changeset
  1910
  Most of the used notions can also be looked up in
f07b8fb99818 more document structure
haftmann
parents: 70144
diff changeset
  1911
    \<^item> \<^url>\<open>http://www.mathworld.com\<close> by Eric Weisstein et. al.
f07b8fb99818 more document structure
haftmann
parents: 70144
diff changeset
  1912
    \<^item> \<^emph>\<open>Algebra I\<close> by van der Waerden, Springer
f07b8fb99818 more document structure
haftmann
parents: 70144
diff changeset
  1913
\<close>
66807
c3631f32dfeb tuned structure
haftmann
parents: 66793
diff changeset
  1914
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
  1915
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
  1916
  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
  1917
  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
  1918
begin
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1919
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1920
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
  1921
  apply (erule (1) mult_right_mono [THEN order_trans])
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1922
  apply (erule (1) mult_left_mono)
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1923
  done
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1924
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1925
lemma mult_mono': "a \<le> b \<Longrightarrow> c \<le> d \<Longrightarrow> 0 \<le> a \<Longrightarrow> 0 \<le> c \<Longrightarrow> a * c \<le> b * d"
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  1926
  by (rule mult_mono) (fast intro: order_trans)+
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1927
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1928
end
21199
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
  1929
74007
df976eefcba0 A few new lemmas and simplifications
paulson <lp15@cam.ac.uk>
parents: 73545
diff changeset
  1930
lemma mono_mult:
df976eefcba0 A few new lemmas and simplifications
paulson <lp15@cam.ac.uk>
parents: 73545
diff changeset
  1931
  fixes a :: "'a::ordered_semiring" 
df976eefcba0 A few new lemmas and simplifications
paulson <lp15@cam.ac.uk>
parents: 73545
diff changeset
  1932
  shows "a \<ge> 0 \<Longrightarrow> mono ((*) a)"
df976eefcba0 A few new lemmas and simplifications
paulson <lp15@cam.ac.uk>
parents: 73545
diff changeset
  1933
  by (simp add: mono_def mult_left_mono)
df976eefcba0 A few new lemmas and simplifications
paulson <lp15@cam.ac.uk>
parents: 73545
diff changeset
  1934
62377
ace69956d018 moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
hoelzl
parents: 62376
diff changeset
  1935
class ordered_semiring_0 = semiring_0 + ordered_semiring
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
  1936
begin
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1937
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  1938
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
  1939
  using mult_left_mono [of 0 b a] by simp
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1940
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1941
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
  1942
  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
  1943
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
  1944
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
  1945
  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
  1946
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  1947
text \<open>Legacy -- use @{thm [source] mult_nonpos_nonneg}.\<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
  1948
lemma mult_nonneg_nonpos2: "0 \<le> a \<Longrightarrow> b \<le> 0 \<Longrightarrow> b * a \<le> 0"
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  1949
  by (drule mult_right_mono [of b 0]) auto
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1950
62378
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  1951
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
  1952
  by (auto simp add: mult_nonneg_nonpos mult_nonneg_nonpos2)
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1953
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1954
end
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1955
82595
c0587d661ea8 Tweaking the ordered semiring type classes
paulson <lp15@cam.ac.uk>
parents: 82593
diff changeset
  1956
class zero_less_one = order + zero + one +
c0587d661ea8 Tweaking the ordered semiring type classes
paulson <lp15@cam.ac.uk>
parents: 82593
diff changeset
  1957
  assumes zero_less_one [simp]: "0 < 1"
c0587d661ea8 Tweaking the ordered semiring type classes
paulson <lp15@cam.ac.uk>
parents: 82593
diff changeset
  1958
begin
c0587d661ea8 Tweaking the ordered semiring type classes
paulson <lp15@cam.ac.uk>
parents: 82593
diff changeset
  1959
c0587d661ea8 Tweaking the ordered semiring type classes
paulson <lp15@cam.ac.uk>
parents: 82593
diff changeset
  1960
subclass zero_neq_one
c0587d661ea8 Tweaking the ordered semiring type classes
paulson <lp15@cam.ac.uk>
parents: 82593
diff changeset
  1961
  by standard (simp add: less_imp_neq)
c0587d661ea8 Tweaking the ordered semiring type classes
paulson <lp15@cam.ac.uk>
parents: 82593
diff changeset
  1962
c0587d661ea8 Tweaking the ordered semiring type classes
paulson <lp15@cam.ac.uk>
parents: 82593
diff changeset
  1963
lemma zero_le_one [simp]:
c0587d661ea8 Tweaking the ordered semiring type classes
paulson <lp15@cam.ac.uk>
parents: 82593
diff changeset
  1964
  \<open>0 \<le> 1\<close> by (rule less_imp_le) simp
c0587d661ea8 Tweaking the ordered semiring type classes
paulson <lp15@cam.ac.uk>
parents: 82593
diff changeset
  1965
c0587d661ea8 Tweaking the ordered semiring type classes
paulson <lp15@cam.ac.uk>
parents: 82593
diff changeset
  1966
end
c0587d661ea8 Tweaking the ordered semiring type classes
paulson <lp15@cam.ac.uk>
parents: 82593
diff changeset
  1967
c0587d661ea8 Tweaking the ordered semiring type classes
paulson <lp15@cam.ac.uk>
parents: 82593
diff changeset
  1968
class ordered_semiring_1 = ordered_semiring_0 + semiring_1 + zero_less_one
82593
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  1969
begin
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  1970
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  1971
lemma convex_bound_le:
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  1972
  assumes "x \<le> a" and "y \<le> a" and "0 \<le> u" and "0 \<le> v" and "u + v = 1"
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  1973
  shows "u * x + v * y \<le> a"
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  1974
proof-
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  1975
  from assms have "u * x + v * y \<le> u * a + v * a"
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  1976
    by (simp add: add_mono mult_left_mono)
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  1977
  with assms show ?thesis
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  1978
    unfolding distrib_right[symmetric] by simp
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  1979
qed
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  1980
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  1981
end
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  1982
62377
ace69956d018 moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
hoelzl
parents: 62376
diff changeset
  1983
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
  1984
begin
ace69956d018 moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
hoelzl
parents: 62376
diff changeset
  1985
ace69956d018 moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
hoelzl
parents: 62376
diff changeset
  1986
subclass semiring_0_cancel ..
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  1987
62377
ace69956d018 moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
hoelzl
parents: 62376
diff changeset
  1988
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
  1989
82597
328de89f20f9 More type class things
paulson <lp15@cam.ac.uk>
parents: 82595
diff changeset
  1990
subclass ordered_cancel_ab_semigroup_add ..
328de89f20f9 More type class things
paulson <lp15@cam.ac.uk>
parents: 82595
diff changeset
  1991
62377
ace69956d018 moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
hoelzl
parents: 62376
diff changeset
  1992
end
ace69956d018 moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
hoelzl
parents: 62376
diff changeset
  1993
38642
8fa437809c67 dropped type classes mult_mono and mult_mono1; tuned names of technical rule duplicates
haftmann
parents: 37767
diff changeset
  1994
class linordered_semiring = ordered_semiring + linordered_cancel_ab_semigroup_add
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
  1995
begin
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1996
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1997
subclass ordered_cancel_semiring ..
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1998
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
  1999
subclass ordered_cancel_comm_monoid_add ..
25304
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  2000
63456
3365c8ec67bd sharing simp rules between ordered monoids and rings
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 63359
diff changeset
  2001
subclass ordered_ab_semigroup_monoid_add_imp_le ..
3365c8ec67bd sharing simp rules between ordered monoids and rings
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 63359
diff changeset
  2002
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2003
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
  2004
  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
  2005
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2006
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
  2007
  by (force simp add: mult_right_mono not_le [symmetric])
23521
195fe3fe2831 added ordered_ring and ordered_semiring
obua
parents: 23496
diff changeset
  2008
25186
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
  2009
end
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
  2010
82593
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2011
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2012
class ordered_semiring_strict = semiring + comm_monoid_add + ordered_cancel_ab_semigroup_add +
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2013
  assumes mult_strict_left_mono: "a < b \<Longrightarrow> 0 < c \<Longrightarrow> c * a < c * b"
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2014
  assumes mult_strict_right_mono: "a < b \<Longrightarrow> 0 < c \<Longrightarrow> a * c < b * c"
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2015
begin
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2016
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2017
subclass semiring_0_cancel ..
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2018
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2019
subclass ordered_semiring
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2020
proof
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2021
  fix a b c :: 'a
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2022
  assume \<section>: "a \<le> b" "0 \<le> c"
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2023
  thus "c * a \<le> c * b"
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2024
    unfolding le_less
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2025
    using mult_strict_left_mono by (cases "c = 0") auto
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2026
  show "a * c \<le> b * c"
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2027
    using \<section> by (force simp: le_less intro: mult_strict_right_mono dest: sym)
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2028
qed
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2029
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2030
lemma mult_pos_pos[simp]: "0 < a \<Longrightarrow> 0 < b \<Longrightarrow> 0 < a * b"
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2031
  using mult_strict_left_mono [of 0 b a] by simp
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2032
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2033
lemma mult_pos_neg: "0 < a \<Longrightarrow> b < 0 \<Longrightarrow> a * b < 0"
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2034
  using mult_strict_left_mono [of b 0 a] by simp
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2035
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2036
lemma mult_neg_pos: "a < 0 \<Longrightarrow> 0 < b \<Longrightarrow> a * b < 0"
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2037
  using mult_strict_right_mono [of a 0 b] by simp
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2038
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2039
text \<open>Strict monotonicity in both arguments\<close>
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2040
lemma mult_strict_mono:
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2041
  assumes "a < b" "c < d" "0 < b" "0 \<le> c"
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2042
  shows "a * c < b * d"
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2043
proof-
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2044
  have "a * c \<le> b * c"
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2045
    using assms by (intro mult_right_mono) auto
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2046
  also have "... < b * d"
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2047
    using assms by (intro mult_strict_left_mono) auto
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2048
  finally show ?thesis .
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2049
qed
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2050
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2051
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2052
text \<open>This weaker variant has more natural premises\<close>
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2053
lemma mult_strict_mono':
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2054
  assumes "a < b" and "c < d" and "0 \<le> a" and "0 \<le> c"
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2055
  shows "a * c < b * d"
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2056
  using assms by (intro mult_strict_mono) auto
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2057
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2058
lemma mult_less_le_imp_less:
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2059
  assumes "a < b" "c \<le> d" "0 \<le> a" "0 < c"
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2060
  shows "a * c < b * d"
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2061
proof-
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2062
  have "a * c < b * c"
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2063
    using assms by (intro mult_strict_right_mono) auto
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2064
  also have "... \<le> b * d"
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2065
    using assms by (intro mult_left_mono) auto
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2066
  finally show ?thesis .
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2067
qed
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2068
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2069
lemma mult_le_less_imp_less:
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2070
  assumes "a \<le> b" and "c < d" and "0 < a" and "0 \<le> c"
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2071
  shows "a * c < b * d"
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2072
proof-
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2073
  have "a * c \<le> b * c"
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2074
    using assms by (intro mult_right_mono) auto
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2075
  also have "... < b * d"
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2076
    using assms by (intro mult_strict_left_mono) auto
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2077
  finally show ?thesis .
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2078
qed
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2079
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2080
end
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2081
66937
a1a4a5e2933a rule out pathologic instances
haftmann
parents: 66816
diff changeset
  2082
class linordered_semiring_1 = linordered_semiring + semiring_1 + zero_less_one
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
  2083
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
  2084
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
  2085
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
  2086
  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
  2087
  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
  2088
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
  2089
  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
  2090
    by (simp add: add_mono mult_left_mono)
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2091
  with assms show ?thesis
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2092
    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
  2093
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
  2094
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
  2095
end
82593
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2096
                                        
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2097
subclass (in linordered_semiring_1) ordered_semiring_1 ..
35043
07dbdf60d5ad dropped accidental duplication of "lin" prefix from cs. 108662d50512
haftmann
parents: 35032
diff changeset
  2098
07dbdf60d5ad dropped accidental duplication of "lin" prefix from cs. 108662d50512
haftmann
parents: 35032
diff changeset
  2099
class linordered_semiring_strict = semiring + comm_monoid_add + linordered_cancel_ab_semigroup_add +
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
  2100
  assumes mult_strict_left_mono: "a < b \<Longrightarrow> 0 < c \<Longrightarrow> c * a < c * b"
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
  2101
  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
  2102
begin
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14334
diff changeset
  2103
27516
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
  2104
subclass semiring_0_cancel ..
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
  2105
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  2106
subclass linordered_semiring
28823
dcbef866c9e2 tuned unfold_locales invocation
haftmann
parents: 28559
diff changeset
  2107
proof
23550
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
  2108
  fix a b c :: 'a
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  2109
  assume *: "a \<le> b" "0 \<le> c"
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  2110
  then show "c * a \<le> c * b"
25186
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
  2111
    unfolding le_less
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
  2112
    using mult_strict_left_mono by (cases "c = 0") auto
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  2113
  from * show "a * c \<le> b * c"
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
  2114
    unfolding le_less
25186
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
  2115
    using mult_strict_right_mono by (cases "c = 0") auto
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
  2116
qed
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
  2117
82593
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2118
subclass (in linordered_semiring_strict) ordered_semiring_strict
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2119
proof qed (auto simp: mult_strict_left_mono mult_strict_right_mono)
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2120
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2121
lemma mult_left_le_imp_le: "c * a \<le> c * b \<Longrightarrow> 0 < c \<Longrightarrow> a \<le> b"
82593
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2122
  by (auto simp add: mult_strict_left_mono simp flip: not_less)
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
  2123
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2124
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
  2125
  by (auto simp add: mult_strict_right_mono not_less [symmetric])
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  2126
71697
34ff9ca387c0 fixed more nasty proofs
paulson <lp15@cam.ac.uk>
parents: 71398
diff changeset
  2127
lemma zero_less_mult_pos: 
34ff9ca387c0 fixed more nasty proofs
paulson <lp15@cam.ac.uk>
parents: 71398
diff changeset
  2128
  assumes "0 < a * b" "0 < a" shows "0 < b"
34ff9ca387c0 fixed more nasty proofs
paulson <lp15@cam.ac.uk>
parents: 71398
diff changeset
  2129
proof (cases "b \<le> 0")
34ff9ca387c0 fixed more nasty proofs
paulson <lp15@cam.ac.uk>
parents: 71398
diff changeset
  2130
  case True
34ff9ca387c0 fixed more nasty proofs
paulson <lp15@cam.ac.uk>
parents: 71398
diff changeset
  2131
  then show ?thesis
34ff9ca387c0 fixed more nasty proofs
paulson <lp15@cam.ac.uk>
parents: 71398
diff changeset
  2132
    using assms by (auto simp: le_less dest: less_not_sym mult_pos_neg [of a b])
34ff9ca387c0 fixed more nasty proofs
paulson <lp15@cam.ac.uk>
parents: 71398
diff changeset
  2133
qed (auto simp add: le_less not_less)
34ff9ca387c0 fixed more nasty proofs
paulson <lp15@cam.ac.uk>
parents: 71398
diff changeset
  2134
34ff9ca387c0 fixed more nasty proofs
paulson <lp15@cam.ac.uk>
parents: 71398
diff changeset
  2135
lemma zero_less_mult_pos2: 
34ff9ca387c0 fixed more nasty proofs
paulson <lp15@cam.ac.uk>
parents: 71398
diff changeset
  2136
  assumes "0 < b * a" "0 < a" shows "0 < b"
34ff9ca387c0 fixed more nasty proofs
paulson <lp15@cam.ac.uk>
parents: 71398
diff changeset
  2137
proof (cases "b \<le> 0")
34ff9ca387c0 fixed more nasty proofs
paulson <lp15@cam.ac.uk>
parents: 71398
diff changeset
  2138
  case True
34ff9ca387c0 fixed more nasty proofs
paulson <lp15@cam.ac.uk>
parents: 71398
diff changeset
  2139
  then show ?thesis
82593
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2140
    using assms by (auto simp: le_less dest: less_not_sym mult_neg_pos)
71697
34ff9ca387c0 fixed more nasty proofs
paulson <lp15@cam.ac.uk>
parents: 71398
diff changeset
  2141
qed (auto simp add: le_less not_less)
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2142
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  2143
end
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  2144
66937
a1a4a5e2933a rule out pathologic instances
haftmann
parents: 66816
diff changeset
  2145
class linordered_semiring_1_strict = linordered_semiring_strict + semiring_1 + zero_less_one
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
  2146
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
  2147
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
  2148
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
  2149
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
  2150
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
  2151
  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
  2152
  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
  2153
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
  2154
  from assms have "u * x + v * y < u * a + v * a"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2155
    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
  2156
  with assms show ?thesis
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2157
    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
  2158
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
  2159
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
  2160
end
33319
74f0dcc0b5fb moved algebraic classes to Ring_and_Field
haftmann
parents: 32960
diff changeset
  2161
82595
c0587d661ea8 Tweaking the ordered semiring type classes
paulson <lp15@cam.ac.uk>
parents: 82593
diff changeset
  2162
class ordered_semiring_1_strict = ordered_semiring_strict + semiring_1 + zero_less_one
82593
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2163
  \<comment> \<open>analogous to \<^class>\<open>linordered_semiring_1_strict\<close> not requiring a total order\<close>
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2164
begin
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2165
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2166
subclass ordered_semiring_1 ..
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2167
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2168
lemma convex_bound_lt:
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2169
  assumes "x < a" and "y < a" and "0 \<le> u" and "0 \<le> v" and "u + v = 1"
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2170
  shows "u * x + v * y < a"
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2171
proof -
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2172
  from assms have "u * x + v * y < u * a + v * a"
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2173
    by (cases "u = 0") (auto intro!: add_less_le_mono mult_strict_left_mono mult_left_mono)
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2174
  with assms show ?thesis
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2175
    unfolding distrib_right[symmetric] by simp
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2176
qed
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2177
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2178
end
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2179
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2180
subclass (in linordered_semiring_1_strict) ordered_semiring_1_strict ..
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2181
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
  2182
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
  2183
  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
  2184
begin
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
  2185
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  2186
subclass ordered_semiring
28823
dcbef866c9e2 tuned unfold_locales invocation
haftmann
parents: 28559
diff changeset
  2187
proof
21199
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
  2188
  fix a b c :: 'a
23550
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
  2189
  assume "a \<le> b" "0 \<le> c"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2190
  then show "c * a \<le> c * b" by (rule comm_mult_left_mono)
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2191
  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
  2192
qed
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
  2193
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
  2194
end
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
  2195
38642
8fa437809c67 dropped type classes mult_mono and mult_mono1; tuned names of technical rule duplicates
haftmann
parents: 37767
diff changeset
  2196
class ordered_cancel_comm_semiring = ordered_comm_semiring + cancel_comm_monoid_add
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
  2197
begin
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
  2198
38642
8fa437809c67 dropped type classes mult_mono and mult_mono1; tuned names of technical rule duplicates
haftmann
parents: 37767
diff changeset
  2199
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
  2200
subclass ordered_comm_semiring ..
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  2201
subclass ordered_cancel_semiring ..
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
  2202
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
  2203
end
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
  2204
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  2205
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
  2206
  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
  2207
begin
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
  2208
35043
07dbdf60d5ad dropped accidental duplication of "lin" prefix from cs. 108662d50512
haftmann
parents: 35032
diff changeset
  2209
subclass linordered_semiring_strict
28823
dcbef866c9e2 tuned unfold_locales invocation
haftmann
parents: 28559
diff changeset
  2210
proof
23550
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
  2211
  fix a b c :: 'a
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
  2212
  assume "a < b" "0 < c"
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  2213
  then show "c * a < c * b"
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  2214
    by (rule comm_mult_strict_left_mono)
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  2215
  then show "a * c < b * c"
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  2216
    by (simp only: mult.commute)
23550
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
  2217
qed
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
  2218
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  2219
subclass ordered_cancel_comm_semiring
28823
dcbef866c9e2 tuned unfold_locales invocation
haftmann
parents: 28559
diff changeset
  2220
proof
23550
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
  2221
  fix a b c :: 'a
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
  2222
  assume "a \<le> b" "0 \<le> c"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2223
  then show "c * a \<le> c * b"
25186
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
  2224
    unfolding le_less
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  2225
    using mult_strict_left_mono by (cases "c = 0") auto
23550
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
  2226
qed
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
  2227
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
  2228
end
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  2229
82593
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2230
class ordered_comm_semiring_strict = comm_semiring_0 + ordered_cancel_ab_semigroup_add +
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2231
  \<comment> \<open>analogous to \<^class>\<open>linordered_comm_semiring_strict\<close> not requiring a total order\<close>
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2232
  assumes comm_mult_strict_left_mono: "a < b \<Longrightarrow> 0 < c \<Longrightarrow> c * a < c * b"
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2233
begin
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2234
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2235
subclass ordered_semiring_strict
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2236
proof
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2237
  fix a b c :: 'a
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2238
  assume "a < b" and "0 < c"
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2239
  thus "c * a < c * b"
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2240
    by (rule comm_mult_strict_left_mono)
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2241
  thus "a * c < b * c"
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2242
    by (simp only: mult.commute)
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2243
qed
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2244
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2245
subclass ordered_cancel_comm_semiring
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2246
proof
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2247
  fix a b c :: 'a
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2248
  assume "a \<le> b" and "0 \<le> c"
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2249
  thus "c * a \<le> c * b"
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2250
    unfolding le_less
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2251
    using mult_strict_left_mono by (cases "c = 0") auto
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2252
qed
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2253
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2254
end
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2255
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2256
subclass (in linordered_comm_semiring_strict) ordered_comm_semiring_strict
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2257
proof qed (simp add: local.mult_strict_left_mono)
88043331f166 New type classes unboundedx_dense_order, ordered_semiring_1, ordered_semiring_strict and related material from Complex_Bounded_Operators
paulson <lp15@cam.ac.uk>
parents: 82518
diff changeset
  2258
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
  2259
class ordered_ring = ring + ordered_cancel_semiring
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
  2260
begin
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  2261
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  2262
subclass ordered_ab_group_add ..
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  2263
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2264
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
  2265
  by (simp add: algebra_simps)
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  2266
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2267
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
  2268
  by (simp add: algebra_simps)
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  2269
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2270
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
  2271
  by (simp add: algebra_simps)
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  2272
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2273
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
  2274
  by (simp add: algebra_simps)
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  2275
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2276
lemma mult_left_mono_neg: "b \<le> a \<Longrightarrow> c \<le> 0 \<Longrightarrow> c * a \<le> c * b"
71697
34ff9ca387c0 fixed more nasty proofs
paulson <lp15@cam.ac.uk>
parents: 71398
diff changeset
  2277
  by (auto dest: mult_left_mono [of _ _ "- c"])
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  2278
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2279
lemma mult_right_mono_neg: "b \<le> a \<Longrightarrow> c \<le> 0 \<Longrightarrow> a * c \<le> b * c"
71697
34ff9ca387c0 fixed more nasty proofs
paulson <lp15@cam.ac.uk>
parents: 71398
diff changeset
  2280
  by (auto dest: mult_right_mono [of _ _ "- c"])
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  2281
30692
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
  2282
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
  2283
  using mult_right_mono_neg [of a 0 b] by simp
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  2284
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2285
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
  2286
  by (auto simp add: mult_nonpos_nonpos)
25186
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
  2287
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
  2288
end
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  2289
64290
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
  2290
class abs_if = minus + uminus + ord + zero + abs +
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
  2291
  assumes abs_if: "\<bar>a\<bar> = (if a < 0 then - a else a)"
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
  2292
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  2293
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
  2294
begin
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  2295
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  2296
subclass ordered_ring ..
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  2297
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  2298
subclass ordered_ab_group_add_abs
28823
dcbef866c9e2 tuned unfold_locales invocation
haftmann
parents: 28559
diff changeset
  2299
proof
25304
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  2300
  fix a b
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  2301
  show "\<bar>a + b\<bar> \<le> \<bar>a\<bar> + \<bar>b\<bar>"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2302
    by (auto simp add: abs_if not_le not_less algebra_simps
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2303
        simp del: add.commute dest: add_neg_neg add_nonneg_nonneg)
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  2304
qed (auto simp: abs_if)
25304
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  2305
35631
0b8a5fd339ab generalize some lemmas from class linordered_ring_strict to linordered_ring
huffman
parents: 35302
diff changeset
  2306
lemma zero_le_square [simp]: "0 \<le> a * a"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2307
  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
  2308
0b8a5fd339ab generalize some lemmas from class linordered_ring_strict to linordered_ring
huffman
parents: 35302
diff changeset
  2309
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
  2310
  by (simp add: not_less)
0b8a5fd339ab generalize some lemmas from class linordered_ring_strict to linordered_ring
huffman
parents: 35302
diff changeset
  2311
61944
5d06ecfdb472 prefer symbols for "abs";
wenzelm
parents: 61799
diff changeset
  2312
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
  2313
  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
  2314
64848
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  2315
lemma abs_eq_iff':
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  2316
  "\<bar>a\<bar> = b \<longleftrightarrow> b \<ge> 0 \<and> (a = b \<or> a = - b)"
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  2317
  by (cases "a \<ge> 0") auto
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  2318
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  2319
lemma eq_abs_iff':
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  2320
  "a = \<bar>b\<bar> \<longleftrightarrow> a \<ge> 0 \<and> (b = a \<or> b = - a)"
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  2321
  using abs_eq_iff' [of b a] by auto
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64713
diff changeset
  2322
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2323
lemma sum_squares_ge_zero: "0 \<le> x * x + y * y"
62347
2230b7047376 generalized some lemmas;
haftmann
parents: 61944
diff changeset
  2324
  by (intro add_nonneg_nonneg zero_le_square)
2230b7047376 generalized some lemmas;
haftmann
parents: 61944
diff changeset
  2325
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2326
lemma not_sum_squares_lt_zero: "\<not> x * x + y * y < 0"
62347
2230b7047376 generalized some lemmas;
haftmann
parents: 61944
diff changeset
  2327
  by (simp add: not_less sum_squares_ge_zero)
2230b7047376 generalized some lemmas;
haftmann
parents: 61944
diff changeset
  2328
25304
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  2329
end
23521
195fe3fe2831 added ordered_ring and ordered_semiring
obua
parents: 23496
diff changeset
  2330
35043
07dbdf60d5ad dropped accidental duplication of "lin" prefix from cs. 108662d50512
haftmann
parents: 35032
diff changeset
  2331
class linordered_ring_strict = ring + linordered_semiring_strict
25304
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  2332
  + ordered_ab_group_add + abs_if
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  2333
begin
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
  2334
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  2335
subclass linordered_ring ..
25304
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  2336
30692
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
  2337
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
  2338
  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
  2339
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
  2340
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
  2341
  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
  2342
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
  2343
lemma mult_neg_neg: "a < 0 \<Longrightarrow> b < 0 \<Longrightarrow> 0 < a * b"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2344
  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
  2345
25917
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  2346
subclass ring_no_zero_divisors
28823
dcbef866c9e2 tuned unfold_locales invocation
haftmann
parents: 28559
diff changeset
  2347
proof
25917
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  2348
  fix a b
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2349
  assume "a \<noteq> 0"
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  2350
  then have a: "a < 0 \<or> 0 < a" by (simp add: neq_iff)
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2351
  assume "b \<noteq> 0"
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  2352
  then have b: "b < 0 \<or> 0 < b" by (simp add: neq_iff)
25917
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  2353
  have "a * b < 0 \<or> 0 < a * b"
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  2354
  proof (cases "a < 0")
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  2355
    case True
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2356
    show ?thesis
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2357
    proof (cases "b < 0")
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2358
      case True
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  2359
      with \<open>a < 0\<close> show ?thesis by (auto dest: mult_neg_neg)
25917
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  2360
    next
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2361
      case False
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  2362
      with b have "0 < b" by auto
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  2363
      with \<open>a < 0\<close> show ?thesis by (auto dest: mult_strict_right_mono)
25917
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  2364
    qed
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  2365
  next
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2366
    case False
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  2367
    with a have "0 < a" by auto
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2368
    show ?thesis
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2369
    proof (cases "b < 0")
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2370
      case True
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  2371
      with \<open>0 < a\<close> show ?thesis
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2372
        by (auto dest: mult_strict_right_mono_neg)
25917
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  2373
    next
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2374
      case False
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  2375
      with b have "0 < b" by auto
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  2376
      with \<open>0 < a\<close> show ?thesis by auto
25917
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  2377
    qed
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  2378
  qed
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2379
  then show "a * b \<noteq> 0"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2380
    by (simp add: neq_iff)
25917
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  2381
qed
25304
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  2382
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70347
diff changeset
  2383
lemma zero_less_mult_iff [algebra_split_simps, field_split_simps]:
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70347
diff changeset
  2384
  "0 < a * b \<longleftrightarrow> 0 < a \<and> 0 < b \<or> a < 0 \<and> b < 0"
56480
093ea91498e6 field_simps: better support for negation and division, and power
hoelzl
parents: 56217
diff changeset
  2385
  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
  2386
     (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
  2387
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70347
diff changeset
  2388
lemma zero_le_mult_iff [algebra_split_simps, field_split_simps]:
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70347
diff changeset
  2389
  "0 \<le> a * b \<longleftrightarrow> 0 \<le> a \<and> 0 \<le> b \<or> a \<le> 0 \<and> b \<le> 0"
56480
093ea91498e6 field_simps: better support for negation and division, and power
hoelzl
parents: 56217
diff changeset
  2390
  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
  2391
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70347
diff changeset
  2392
lemma mult_less_0_iff [algebra_split_simps, field_split_simps]:
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70347
diff changeset
  2393
  "a * b < 0 \<longleftrightarrow> 0 < a \<and> b < 0 \<or> a < 0 \<and> 0 < b"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2394
  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
  2395
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70347
diff changeset
  2396
lemma mult_le_0_iff [algebra_split_simps, field_split_simps]:
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70347
diff changeset
  2397
  "a * b \<le> 0 \<longleftrightarrow> 0 \<le> a \<and> b \<le> 0 \<or> a \<le> 0 \<and> 0 \<le> b"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2398
  using zero_le_mult_iff [of "- a" b] by auto
25917
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  2399
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2400
text \<open>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68253
diff changeset
  2401
  Cancellation laws for \<^term>\<open>c * a < c * b\<close> and \<^term>\<open>a * c < b * c\<close>,
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2402
  also with the relations \<open>\<le>\<close> and equality.
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2403
\<close>
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  2404
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2405
text \<open>
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2406
  These ``disjunction'' versions produce two cases when the comparison is
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2407
  an assumption, but effectively four when the comparison is a goal.
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2408
\<close>
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  2409
71699
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2410
lemma mult_less_cancel_right_disj: "a * c < b * c \<longleftrightarrow> 0 < c \<and> a < b \<or> c < 0 \<and> b < a"
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2411
proof (cases "c = 0")
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2412
  case False
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2413
  show ?thesis (is "?lhs \<longleftrightarrow> ?rhs")
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2414
  proof
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2415
    assume ?lhs
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2416
    then have "c < 0 \<Longrightarrow> b < a" "c > 0 \<Longrightarrow> b > a"
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2417
      by (auto simp flip: not_le intro: mult_right_mono mult_right_mono_neg)
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2418
    with False show ?rhs 
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2419
      by (auto simp add: neq_iff)
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2420
  next
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2421
    assume ?rhs
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2422
    with False show ?lhs 
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2423
      by (auto simp add: mult_strict_right_mono mult_strict_right_mono_neg)
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2424
  qed
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2425
qed auto
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2426
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2427
lemma mult_less_cancel_left_disj: "c * a < c * b \<longleftrightarrow> 0 < c \<and> a < b \<or> c < 0 \<and> b < a"
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2428
proof (cases "c = 0")
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2429
  case False
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2430
  show ?thesis (is "?lhs \<longleftrightarrow> ?rhs")
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2431
  proof
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2432
    assume ?lhs
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2433
    then have "c < 0 \<Longrightarrow> b < a" "c > 0 \<Longrightarrow> b > a"
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2434
      by (auto simp flip: not_le intro: mult_left_mono mult_left_mono_neg)
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2435
    with False show ?rhs 
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2436
      by (auto simp add: neq_iff)
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2437
  next
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2438
    assume ?rhs
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2439
    with False show ?lhs 
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2440
      by (auto simp add: mult_strict_left_mono mult_strict_left_mono_neg)
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2441
  qed
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2442
qed auto
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  2443
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2444
text \<open>
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2445
  The ``conjunction of implication'' lemmas produce two cases when the
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2446
  comparison is a goal, but give four when the comparison is an assumption.
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2447
\<close>
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  2448
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2449
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
  2450
  using mult_less_cancel_right_disj [of a c b] by auto
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  2451
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2452
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
  2453
  using mult_less_cancel_left_disj [of c a b] by auto
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  2454
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2455
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
  2456
  by (simp add: not_less [symmetric] mult_less_cancel_right_disj)
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  2457
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2458
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
  2459
  by (simp add: not_less [symmetric] mult_less_cancel_left_disj)
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  2460
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2461
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
  2462
  by (auto simp: mult_le_cancel_left)
30649
57753e0ec1d4 1. New cancellation simprocs for common factors in inequations
nipkow
parents: 30242
diff changeset
  2463
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2464
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
  2465
  by (auto simp: mult_le_cancel_left)
30649
57753e0ec1d4 1. New cancellation simprocs for common factors in inequations
nipkow
parents: 30242
diff changeset
  2466
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2467
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
  2468
  by (auto simp: mult_less_cancel_left)
30649
57753e0ec1d4 1. New cancellation simprocs for common factors in inequations
nipkow
parents: 30242
diff changeset
  2469
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2470
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
  2471
  by (auto simp: mult_less_cancel_left)
30649
57753e0ec1d4 1. New cancellation simprocs for common factors in inequations
nipkow
parents: 30242
diff changeset
  2472
79566
f783490c6c99 A small number of new lemmas
paulson <lp15@cam.ac.uk>
parents: 79541
diff changeset
  2473
lemma mult_le_cancel_right_pos: "0 < c \<Longrightarrow> a * c \<le> b * c \<longleftrightarrow> a \<le> b"
f783490c6c99 A small number of new lemmas
paulson <lp15@cam.ac.uk>
parents: 79541
diff changeset
  2474
  by (auto simp: mult_le_cancel_right)
f783490c6c99 A small number of new lemmas
paulson <lp15@cam.ac.uk>
parents: 79541
diff changeset
  2475
f783490c6c99 A small number of new lemmas
paulson <lp15@cam.ac.uk>
parents: 79541
diff changeset
  2476
lemma mult_le_cancel_right_neg: "c < 0 \<Longrightarrow> a * c \<le> b * c \<longleftrightarrow> b \<le> a"
f783490c6c99 A small number of new lemmas
paulson <lp15@cam.ac.uk>
parents: 79541
diff changeset
  2477
  by (auto simp: mult_le_cancel_right)
f783490c6c99 A small number of new lemmas
paulson <lp15@cam.ac.uk>
parents: 79541
diff changeset
  2478
f783490c6c99 A small number of new lemmas
paulson <lp15@cam.ac.uk>
parents: 79541
diff changeset
  2479
lemma mult_less_cancel_right_pos: "0 < c \<Longrightarrow> a * c < b * c \<longleftrightarrow> a < b"
f783490c6c99 A small number of new lemmas
paulson <lp15@cam.ac.uk>
parents: 79541
diff changeset
  2480
  by (auto simp: mult_less_cancel_right)
f783490c6c99 A small number of new lemmas
paulson <lp15@cam.ac.uk>
parents: 79541
diff changeset
  2481
f783490c6c99 A small number of new lemmas
paulson <lp15@cam.ac.uk>
parents: 79541
diff changeset
  2482
lemma mult_less_cancel_right_neg: "c < 0 \<Longrightarrow> a * c < b * c \<longleftrightarrow> b < a"
f783490c6c99 A small number of new lemmas
paulson <lp15@cam.ac.uk>
parents: 79541
diff changeset
  2483
  by (auto simp: mult_less_cancel_right)
f783490c6c99 A small number of new lemmas
paulson <lp15@cam.ac.uk>
parents: 79541
diff changeset
  2484
25917
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  2485
end
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
  2486
30692
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
  2487
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
  2488
  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
  2489
  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
  2490
  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
  2491
  mult_neg_pos mult_neg_neg
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  2492
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  2493
class ordered_comm_ring = comm_ring + ordered_comm_semiring
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
  2494
begin
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  2495
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  2496
subclass ordered_ring ..
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  2497
subclass ordered_cancel_comm_semiring ..
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  2498
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
  2499
end
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  2500
67689
2c38ffd6ec71 type class linordered_nonzero_semiring has new axiom to guarantee characteristic 0
paulson <lp15@cam.ac.uk>
parents: 67234
diff changeset
  2501
class linordered_nonzero_semiring = ordered_comm_semiring + monoid_mult + linorder + zero_less_one +
2c38ffd6ec71 type class linordered_nonzero_semiring has new axiom to guarantee characteristic 0
paulson <lp15@cam.ac.uk>
parents: 67234
diff changeset
  2502
  assumes add_mono1: "a < b \<Longrightarrow> a + 1 < b + 1"
62378
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  2503
begin
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  2504
82595
c0587d661ea8 Tweaking the ordered semiring type classes
paulson <lp15@cam.ac.uk>
parents: 82593
diff changeset
  2505
subclass zero_neq_one ..
62378
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  2506
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  2507
subclass comm_semiring_1
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2508
  by standard (rule mult_1_left)
62378
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  2509
82595
c0587d661ea8 Tweaking the ordered semiring type classes
paulson <lp15@cam.ac.uk>
parents: 82593
diff changeset
  2510
subclass ordered_semiring_1 ..
62378
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  2511
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  2512
lemma not_one_le_zero [simp]: "\<not> 1 \<le> 0"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2513
  by (simp add: not_le)
62378
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  2514
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  2515
lemma not_one_less_zero [simp]: "\<not> 1 < 0"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2516
  by (simp add: not_less)
62378
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  2517
73535
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
  2518
lemma of_bool_less_eq_iff [simp]:
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
  2519
  \<open>of_bool P \<le> of_bool Q \<longleftrightarrow> (P \<longrightarrow> Q)\<close>
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
  2520
  by auto
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
  2521
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
  2522
lemma of_bool_less_iff [simp]:
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
  2523
  \<open>of_bool P < of_bool Q \<longleftrightarrow> \<not> P \<and> Q\<close>
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
  2524
  by auto
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
  2525
62378
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  2526
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
  2527
  using mult_left_mono[of c 1 a] by simp
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  2528
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  2529
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
  2530
  using mult_mono[of a 1 b 1] by simp
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  2531
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  2532
lemma zero_less_two: "0 < 1 + 1"
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  2533
  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
  2534
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  2535
end
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  2536
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  2537
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
  2538
  assumes le_add_diff_inverse2 [simp]: "b \<le> a \<Longrightarrow> a - b + b = a"
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  2539
begin
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  2540
67689
2c38ffd6ec71 type class linordered_nonzero_semiring has new axiom to guarantee characteristic 0
paulson <lp15@cam.ac.uk>
parents: 67234
diff changeset
  2541
subclass linordered_nonzero_semiring 
2c38ffd6ec71 type class linordered_nonzero_semiring has new axiom to guarantee characteristic 0
paulson <lp15@cam.ac.uk>
parents: 67234
diff changeset
  2542
proof
2c38ffd6ec71 type class linordered_nonzero_semiring has new axiom to guarantee characteristic 0
paulson <lp15@cam.ac.uk>
parents: 67234
diff changeset
  2543
  show "a + 1 < b + 1" if "a < b" for a b
75669
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75543
diff changeset
  2544
  proof (rule ccontr)
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75543
diff changeset
  2545
    assume "\<not> a + 1 < b + 1"
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75543
diff changeset
  2546
    moreover with that have "a + 1 < b + 1"
75455
91c16c5ad3e9 tidied auto / simp with null arguments
paulson <lp15@cam.ac.uk>
parents: 75087
diff changeset
  2547
      by simp
75669
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75543
diff changeset
  2548
    ultimately show False
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 75543
diff changeset
  2549
      by contradiction
67689
2c38ffd6ec71 type class linordered_nonzero_semiring has new axiom to guarantee characteristic 0
paulson <lp15@cam.ac.uk>
parents: 67234
diff changeset
  2550
  qed
2c38ffd6ec71 type class linordered_nonzero_semiring has new axiom to guarantee characteristic 0
paulson <lp15@cam.ac.uk>
parents: 67234
diff changeset
  2551
qed
62378
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  2552
73535
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
  2553
lemma zero_less_eq_of_bool [simp]:
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
  2554
  \<open>0 \<le> of_bool P\<close>
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
  2555
  by simp
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
  2556
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
  2557
lemma zero_less_of_bool_iff [simp]:
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
  2558
  \<open>0 < of_bool P \<longleftrightarrow> P\<close>
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
  2559
  by simp
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
  2560
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
  2561
lemma of_bool_less_eq_one [simp]:
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
  2562
  \<open>of_bool P \<le> 1\<close>
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
  2563
  by simp
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
  2564
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
  2565
lemma of_bool_less_one_iff [simp]:
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
  2566
  \<open>of_bool P < 1 \<longleftrightarrow> \<not> P\<close>
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
  2567
  by simp
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
  2568
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
  2569
lemma of_bool_or_iff [simp]:
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
  2570
  \<open>of_bool (P \<or> Q) = max (of_bool P) (of_bool Q)\<close>
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
  2571
  by (simp add: max_def)
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
  2572
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
  2573
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
  2574
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
  2575
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
  2576
  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
  2577
62378
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62377
diff changeset
  2578
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
  2579
  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
  2580
71699
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2581
lemma add_le_imp_le_diff: 
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2582
  assumes "i + k \<le> n" shows "i \<le> n - k"
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2583
proof -
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2584
  have "n - (i + k) + i + k = n"
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2585
    by (simp add: assms add.assoc)
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2586
  with assms add_implies_diff have "i + k \<le> n - k + k"
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2587
    by fastforce
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2588
  then show ?thesis
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2589
    by simp
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2590
qed
60615
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 60570
diff changeset
  2591
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
  2592
lemma add_le_add_imp_diff_le:
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2593
  assumes 1: "i + k \<le> n"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2594
    and 2: "n \<le> j + k"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2595
  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
  2596
proof -
71699
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2597
  have "n - (i + k) + i + k = n"
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2598
    using 1 by (simp add: add.assoc)
60615
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 60570
diff changeset
  2599
  moreover have "n - k = n - k - i + i"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2600
    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
  2601
  ultimately show ?thesis
71699
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2602
    using 2 add_le_imp_le_diff [of "n-k" k "j + k"]
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71697
diff changeset
  2603
    by (simp add: add.commute diff_diff_add)
60615
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 60570
diff changeset
  2604
qed
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 60570
diff changeset
  2605
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2606
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
  2607
  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
  2608
82600
f62666eea755 two more lemmas from the AFP
paulson <lp15@cam.ac.uk>
parents: 82597
diff changeset
  2609
lemma less_1_mult': 
f62666eea755 two more lemmas from the AFP
paulson <lp15@cam.ac.uk>
parents: 82597
diff changeset
  2610
  shows "1 < a \<Longrightarrow> 1 \<le> b \<Longrightarrow> 1 < a * b"
f62666eea755 two more lemmas from the AFP
paulson <lp15@cam.ac.uk>
parents: 82597
diff changeset
  2611
  by (cases "b=1") (auto simp: le_less less_1_mult)
f62666eea755 two more lemmas from the AFP
paulson <lp15@cam.ac.uk>
parents: 82597
diff changeset
  2612
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  2613
end
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  2614
66937
a1a4a5e2933a rule out pathologic instances
haftmann
parents: 66816
diff changeset
  2615
class linordered_idom = comm_ring_1 + linordered_comm_semiring_strict +
a1a4a5e2933a rule out pathologic instances
haftmann
parents: 66816
diff changeset
  2616
  ordered_ab_group_add + abs_if + sgn +
64290
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
  2617
  assumes sgn_if: "sgn x = (if x = 0 then 0 else if 0 < x then 1 else - 1)"
25917
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  2618
begin
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  2619
35043
07dbdf60d5ad dropped accidental duplication of "lin" prefix from cs. 108662d50512
haftmann
parents: 35032
diff changeset
  2620
subclass linordered_ring_strict ..
66937
a1a4a5e2933a rule out pathologic instances
haftmann
parents: 66816
diff changeset
  2621
a1a4a5e2933a rule out pathologic instances
haftmann
parents: 66816
diff changeset
  2622
subclass linordered_semiring_1_strict
a1a4a5e2933a rule out pathologic instances
haftmann
parents: 66816
diff changeset
  2623
proof
a1a4a5e2933a rule out pathologic instances
haftmann
parents: 66816
diff changeset
  2624
  have "0 \<le> 1 * 1"
a1a4a5e2933a rule out pathologic instances
haftmann
parents: 66816
diff changeset
  2625
    by (fact zero_le_square)
a1a4a5e2933a rule out pathologic instances
haftmann
parents: 66816
diff changeset
  2626
  then show "0 < 1" 
a1a4a5e2933a rule out pathologic instances
haftmann
parents: 66816
diff changeset
  2627
    by (simp add: le_less)
a1a4a5e2933a rule out pathologic instances
haftmann
parents: 66816
diff changeset
  2628
qed
a1a4a5e2933a rule out pathologic instances
haftmann
parents: 66816
diff changeset
  2629
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  2630
subclass ordered_comm_ring ..
27516
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
  2631
subclass idom ..
25917
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  2632
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  2633
subclass linordered_semidom
66937
a1a4a5e2933a rule out pathologic instances
haftmann
parents: 66816
diff changeset
  2634
  by standard simp
25917
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  2635
64290
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
  2636
subclass idom_abs_sgn
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
  2637
  by standard
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
  2638
    (auto simp add: sgn_if abs_if zero_less_mult_iff)
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
  2639
73535
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
  2640
lemma abs_bool_eq [simp]:
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
  2641
  \<open>\<bar>of_bool P\<bar> = of_bool P\<close>
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
  2642
  by simp
0f33c7031ec9 new lemmas
haftmann
parents: 71699
diff changeset
  2643
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  2644
lemma linorder_neqE_linordered_idom:
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2645
  assumes "x \<noteq> y"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2646
  obtains "x < y" | "y < x"
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  2647
  using assms by (rule neqE)
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  2648
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  2649
text \<open>These cancellation simp rules also produce two cases when the comparison is a goal.\<close>
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  2650
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2651
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
  2652
  using mult_le_cancel_right [of 1 c b] by simp
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  2653
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2654
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
  2655
  using mult_le_cancel_right [of a c 1] by simp
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  2656
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2657
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
  2658
  using mult_le_cancel_left [of c 1 b] by simp
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  2659
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2660
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
  2661
  using mult_le_cancel_left [of c a 1] by simp
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  2662
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2663
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
  2664
  using mult_less_cancel_right [of 1 c b] by simp
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  2665
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2666
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
  2667
  using mult_less_cancel_right [of a c 1] by simp
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  2668
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2669
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
  2670
  using mult_less_cancel_left [of c 1 b] by simp
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  2671
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2672
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
  2673
  using mult_less_cancel_left [of c a 1] by simp
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  2674
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2675
lemma sgn_0_0: "sgn a = 0 \<longleftrightarrow> a = 0"
64290
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
  2676
  by (fact sgn_eq_0_iff)
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
  2677
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2678
lemma sgn_1_pos: "sgn a = 1 \<longleftrightarrow> a > 0"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2679
  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
  2680
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2681
lemma sgn_1_neg: "sgn a = - 1 \<longleftrightarrow> a < 0"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2682
  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
  2683
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2684
lemma sgn_pos [simp]: "0 < a \<Longrightarrow> sgn a = 1"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2685
  by (simp only: sgn_1_pos)
29940
83b373f61d41 more default simp rules for sgn
haftmann
parents: 29925
diff changeset
  2686
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2687
lemma sgn_neg [simp]: "a < 0 \<Longrightarrow> sgn a = - 1"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2688
  by (simp only: sgn_1_neg)
29940
83b373f61d41 more default simp rules for sgn
haftmann
parents: 29925
diff changeset
  2689
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  2690
lemma abs_sgn: "\<bar>k\<bar> = k * sgn k"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2691
  unfolding sgn_if abs_if by auto
29700
22faf21db3df added some simp rules
nipkow
parents: 29668
diff changeset
  2692
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2693
lemma sgn_greater [simp]: "0 < sgn a \<longleftrightarrow> 0 < a"
29940
83b373f61d41 more default simp rules for sgn
haftmann
parents: 29925
diff changeset
  2694
  unfolding sgn_if by auto
83b373f61d41 more default simp rules for sgn
haftmann
parents: 29925
diff changeset
  2695
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2696
lemma sgn_less [simp]: "sgn a < 0 \<longleftrightarrow> a < 0"
29940
83b373f61d41 more default simp rules for sgn
haftmann
parents: 29925
diff changeset
  2697
  unfolding sgn_if by auto
83b373f61d41 more default simp rules for sgn
haftmann
parents: 29925
diff changeset
  2698
64239
de5cd9217d4c added lemma
haftmann
parents: 64164
diff changeset
  2699
lemma abs_sgn_eq_1 [simp]:
de5cd9217d4c added lemma
haftmann
parents: 64164
diff changeset
  2700
  "a \<noteq> 0 \<Longrightarrow> \<bar>sgn a\<bar> = 1"
64290
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
  2701
  by simp
64239
de5cd9217d4c added lemma
haftmann
parents: 64164
diff changeset
  2702
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2703
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
  2704
  by (simp add: sgn_if)
2230b7047376 generalized some lemmas;
haftmann
parents: 61944
diff changeset
  2705
64713
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2706
lemma sgn_mult_self_eq [simp]:
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2707
  "sgn a * sgn a = of_bool (a \<noteq> 0)"
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2708
  by (cases "a > 0") simp_all
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2709
75875
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  2710
lemma left_sgn_mult_self_eq [simp]:
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  2711
  \<open>sgn a * (sgn a * b) = of_bool (a \<noteq> 0) * b\<close>
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  2712
  by (simp flip: mult.assoc)
48d032035744 streamlined primitive definitions for integer division
haftmann
parents: 75669
diff changeset
  2713
64713
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2714
lemma abs_mult_self_eq [simp]:
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2715
  "\<bar>a\<bar> * \<bar>a\<bar> = a * a"
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2716
  by (cases "a > 0") simp_all
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2717
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2718
lemma same_sgn_sgn_add:
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2719
  "sgn (a + b) = sgn a" if "sgn b = sgn a"
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2720
proof (cases a 0 rule: linorder_cases)
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2721
  case equal
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2722
  with that show ?thesis
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2723
    by simp
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2724
next
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2725
  case less
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2726
  with that have "b < 0"
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2727
    by (simp add: sgn_1_neg)
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2728
  with \<open>a < 0\<close> have "a + b < 0"
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2729
    by (rule add_neg_neg)
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2730
  with \<open>a < 0\<close> show ?thesis
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2731
    by simp
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2732
next
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2733
  case greater
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2734
  with that have "b > 0"
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2735
    by (simp add: sgn_1_pos)
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2736
  with \<open>a > 0\<close> have "a + b > 0"
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2737
    by (rule add_pos_pos)
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2738
  with \<open>a > 0\<close> show ?thesis
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2739
    by simp
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2740
qed
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2741
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2742
lemma same_sgn_abs_add:
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2743
  "\<bar>a + b\<bar> = \<bar>a\<bar> + \<bar>b\<bar>" if "sgn b = sgn a"
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2744
proof -
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2745
  have "a + b = sgn a * \<bar>a\<bar> + sgn b * \<bar>b\<bar>"
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2746
    by (simp add: sgn_mult_abs)
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2747
  also have "\<dots> = sgn a * (\<bar>a\<bar> + \<bar>b\<bar>)"
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2748
    using that by (simp add: algebra_simps)
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2749
  finally show ?thesis
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2750
    by (auto simp add: abs_mult)
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2751
qed
9638c07283bc more facts on sgn, abs
haftmann
parents: 64592
diff changeset
  2752
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66810
diff changeset
  2753
lemma sgn_not_eq_imp:
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66810
diff changeset
  2754
  "sgn a = - sgn b" if "sgn b \<noteq> sgn a" and "sgn a \<noteq> 0" and "sgn b \<noteq> 0"
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66810
diff changeset
  2755
  using that by (cases "a < 0") (auto simp add: sgn_0_0 sgn_1_pos sgn_1_neg)
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66810
diff changeset
  2756
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  2757
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
  2758
  by (simp add: abs_if)
20a506b8256d lemmas abs_dvd_iff, dvd_abs_iff
huffman
parents: 29940
diff changeset
  2759
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  2760
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
  2761
  by (simp add: abs_if)
29653
ece6a0e9f8af added lemma abs_sng
haftmann
parents: 29465
diff changeset
  2762
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2763
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
  2764
  by (subst abs_dvd_iff [symmetric]) simp
33676
802f5e233e48 moved lemma from Algebra/IntRing to Ring_and_Field
nipkow
parents: 33364
diff changeset
  2765
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2766
text \<open>
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2767
  The following lemmas can be proven in more general structures, but
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2768
  are dangerous as simp rules in absence of @{thm neg_equal_zero},
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2769
  @{thm neg_less_pos}, @{thm neg_less_eq_nonneg}.
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2770
\<close>
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  2771
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2772
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
  2773
  by (fact equation_minus_iff)
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  2774
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2775
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
  2776
  by (subst minus_equation_iff, auto)
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  2777
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2778
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
  2779
  by (fact le_minus_iff)
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  2780
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2781
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
  2782
  by (fact minus_le_iff)
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  2783
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2784
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
  2785
  by (fact less_minus_iff)
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  2786
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2787
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
  2788
  by (fact minus_less_iff)
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  2789
66793
deabce3ccf1f new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents: 65811
diff changeset
  2790
lemma add_less_zeroD:
deabce3ccf1f new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents: 65811
diff changeset
  2791
  shows "x+y < 0 \<Longrightarrow> x<0 \<or> y<0"
deabce3ccf1f new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents: 65811
diff changeset
  2792
  by (auto simp: not_less intro: le_less_trans [of _ "x+y"])
deabce3ccf1f new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents: 65811
diff changeset
  2793
75880
714fad33252e more thorough split rules for div and mod on numerals, tuned split rules setup
haftmann
parents: 75875
diff changeset
  2794
text \<open>
714fad33252e more thorough split rules for div and mod on numerals, tuned split rules setup
haftmann
parents: 75875
diff changeset
  2795
  Is this really better than just rewriting with \<open>abs_if\<close>?
714fad33252e more thorough split rules for div and mod on numerals, tuned split rules setup
haftmann
parents: 75875
diff changeset
  2796
\<close>
714fad33252e more thorough split rules for div and mod on numerals, tuned split rules setup
haftmann
parents: 75875
diff changeset
  2797
lemma abs_split [no_atp]: \<open>P \<bar>a\<bar> \<longleftrightarrow> (0 \<le> a \<longrightarrow> P a) \<and> (a < 0 \<longrightarrow> P (- a))\<close>
714fad33252e more thorough split rules for div and mod on numerals, tuned split rules setup
haftmann
parents: 75875
diff changeset
  2798
  by (force dest: order_less_le_trans simp add: abs_if linorder_not_less)
714fad33252e more thorough split rules for div and mod on numerals, tuned split rules setup
haftmann
parents: 75875
diff changeset
  2799
25917
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  2800
end
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  2801
78935
5e788ff7a489 explicit type class for discrete linordered semidoms
haftmann
parents: 75962
diff changeset
  2802
class discrete_linordered_semidom = linordered_semidom +
5e788ff7a489 explicit type class for discrete linordered semidoms
haftmann
parents: 75962
diff changeset
  2803
  assumes less_iff_succ_less_eq: \<open>a < b \<longleftrightarrow> a + 1 \<le> b\<close>
5e788ff7a489 explicit type class for discrete linordered semidoms
haftmann
parents: 75962
diff changeset
  2804
begin
5e788ff7a489 explicit type class for discrete linordered semidoms
haftmann
parents: 75962
diff changeset
  2805
5e788ff7a489 explicit type class for discrete linordered semidoms
haftmann
parents: 75962
diff changeset
  2806
lemma less_eq_iff_succ_less:
5e788ff7a489 explicit type class for discrete linordered semidoms
haftmann
parents: 75962
diff changeset
  2807
  \<open>a \<le> b \<longleftrightarrow> a < b + 1\<close>
5e788ff7a489 explicit type class for discrete linordered semidoms
haftmann
parents: 75962
diff changeset
  2808
  using less_iff_succ_less_eq [of a \<open>b + 1\<close>] by simp
5e788ff7a489 explicit type class for discrete linordered semidoms
haftmann
parents: 75962
diff changeset
  2809
5e788ff7a489 explicit type class for discrete linordered semidoms
haftmann
parents: 75962
diff changeset
  2810
end
5e788ff7a489 explicit type class for discrete linordered semidoms
haftmann
parents: 75962
diff changeset
  2811
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
  2812
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
  2813
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  2814
context linordered_semidom
25193
e2e1a4b00de3 various localizations
haftmann
parents: 25186
diff changeset
  2815
begin
e2e1a4b00de3 various localizations
haftmann
parents: 25186
diff changeset
  2816
e2e1a4b00de3 various localizations
haftmann
parents: 25186
diff changeset
  2817
lemma less_add_one: "a < a + 1"
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  2818
proof -
25193
e2e1a4b00de3 various localizations
haftmann
parents: 25186
diff changeset
  2819
  have "a + 0 < a + 1"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  2820
    by (blast intro: zero_less_one add_strict_left_mono)
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2821
  then show ?thesis by simp
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  2822
qed
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  2823
25193
e2e1a4b00de3 various localizations
haftmann
parents: 25186
diff changeset
  2824
end
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  2825
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  2826
context linordered_idom
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  2827
begin
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
  2828
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2829
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
  2830
  by (rule mult_left_le)
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  2831
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2832
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
  2833
  by (auto simp add: mult_le_cancel_right2)
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  2834
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  2835
end
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  2836
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60690
diff changeset
  2837
text \<open>Absolute Value\<close>
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  2838
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  2839
context linordered_idom
25304
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  2840
begin
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  2841
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2842
lemma mult_sgn_abs: "sgn x * \<bar>x\<bar> = x"
64290
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
  2843
  by (fact sgn_mult_abs)
25304
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  2844
64290
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
  2845
lemma abs_one: "\<bar>1\<bar> = 1"
fb5c74a58796 suitable logical type class for abs, sgn
haftmann
parents: 64242
diff changeset
  2846
  by (fact abs_1)
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  2847
25304
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  2848
end
24491
8d194c9198ae added constant sgn
nipkow
parents: 24427
diff changeset
  2849
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  2850
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
  2851
  assumes abs_eq_mult:
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  2852
    "(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
  2853
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  2854
context linordered_idom
30961
541bfff659af more localisation
haftmann
parents: 30692
diff changeset
  2855
begin
541bfff659af more localisation
haftmann
parents: 30692
diff changeset
  2856
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2857
subclass ordered_ring_abs
63588
d0e2bad67bd4 misc tuning and modernization;
wenzelm
parents: 63456
diff changeset
  2858
  by standard (auto simp: abs_if not_less mult_less_0_iff)
30961
541bfff659af more localisation
haftmann
parents: 30692
diff changeset
  2859
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  2860
lemma abs_mult_self: "\<bar>a\<bar> * \<bar>a\<bar> = a * a"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66937
diff changeset
  2861
  by (fact abs_mult_self_eq)
30961
541bfff659af more localisation
haftmann
parents: 30692
diff changeset
  2862
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  2863
lemma abs_mult_less:
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2864
  assumes ac: "\<bar>a\<bar> < c"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2865
    and bd: "\<bar>b\<bar> < d"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2866
  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
  2867
proof -
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2868
  from ac have "0 < c"
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2869
    by (blast intro: le_less_trans abs_ge_zero)
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2870
  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
  2871
qed
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  2872
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2873
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
  2874
  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
  2875
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2876
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
  2877
  by (simp add: abs_mult)
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  2878
75543
1910054f8c39 some additional lemmas and a little tidying up
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
  2879
lemma abs_mult_pos': "0 \<le> x \<Longrightarrow> x * \<bar>y\<bar> = \<bar>x * y\<bar>"
1910054f8c39 some additional lemmas and a little tidying up
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
  2880
  by (simp add: abs_mult)
1910054f8c39 some additional lemmas and a little tidying up
paulson <lp15@cam.ac.uk>
parents: 75455
diff changeset
  2881
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2882
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
  2883
  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
  2884
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2885
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
  2886
  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
  2887
62626
de25474ce728 Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents: 62608
diff changeset
  2888
lemma abs_add_one_gt_zero: "0 < 1 + \<bar>x\<bar>"
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2889
  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
  2890
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  2891
end
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  2892
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
  2893
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
  2894
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2895
text \<open>
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2896
  Dioids are the alternative extensions of semirings, a semiring can
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2897
  either be a ring or a dioid but never both.
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2898
\<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
  2899
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
  2900
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
  2901
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
  2902
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
  2903
subclass ordered_semiring
63325
1086d56cde86 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
  2904
  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
  2905
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
  2906
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
  2907
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
  2908
59557
ebd8ecacfba6 establish unique preferred fact names
haftmann
parents: 59555
diff changeset
  2909
hide_fact (open) comm_mult_left_mono comm_mult_strict_left_mono distrib
ebd8ecacfba6 establish unique preferred fact names
haftmann
parents: 59555
diff changeset
  2910
52435
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 51520
diff changeset
  2911
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
  2912
  code_module Rings \<rightharpoonup> (SML) Arith and (OCaml) Arith and (Haskell) Arith
33364
2bd12592c5e8 tuned code setup
haftmann
parents: 33319
diff changeset
  2913
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
  2914
end