# HG changeset patch # User haftmann # Date 1555143107 0 # Node ID f07b8fb99818853cf6dc87342d616ac9a645866b # Parent c925bc0df8270dd03eae56b07c6fbb2837c7e6a5 more document structure diff -r c925bc0df827 -r f07b8fb99818 src/HOL/Rings.thy --- a/src/HOL/Rings.thy Sat Apr 13 08:11:46 2019 +0000 +++ b/src/HOL/Rings.thy Sat Apr 13 08:11:47 2019 +0000 @@ -13,6 +13,8 @@ imports Groups Set Fun begin +subsection \Semirings and rings\ + class semiring = ab_semigroup_add + semigroup_mult + assumes distrib_right[algebra_simps]: "(a + b) * c = a * c + b * c" assumes distrib_left[algebra_simps]: "a * (b + c) = a * b + a * c" @@ -126,7 +128,8 @@ end -text \Abstract divisibility\ + +subsection \Abstract divisibility\ class dvd = times begin @@ -403,6 +406,9 @@ end + +subsection \Towards integral domains\ + class semiring_no_zero_divisors = semiring_0 + assumes no_zero_divisors: "a \ 0 \ b \ 0 \ a * b \ 0" begin @@ -633,17 +639,8 @@ end -text \ - The theory of partially ordered rings is taken from the books: - \<^item> \<^emph>\Lattice Theory\ by Garret Birkhoff, American Mathematical Society, 1979 - \<^item> \<^emph>\Partially Ordered Algebraic Systems\, Pergamon Press, 1963 - - Most of the used notions can also be looked up in - \<^item> \<^url>\http://www.mathworld.com\ by Eric Weisstein et. al. - \<^item> \<^emph>\Algebra I\ by van der Waerden, Springer -\ - -text \Syntactic division operator\ + +subsection \(Partial) Division\ class divide = fixes divide :: "'a \ 'a \ 'a" (infixl "div" 70) @@ -787,6 +784,71 @@ end +text \Integral (semi)domains with cancellation rules\ + +class semidom_divide_cancel = semidom_divide + + assumes div_mult_self1: "b \ 0 \ (a + c * b) div b = c + a div b" + and div_mult_mult1: "c \ 0 \ (c * a) div (c * b) = a div b" +begin + +context + fixes b + assumes "b \ 0" +begin + +lemma div_mult_self2: + "(a + b * c) div b = c + a div b" + using \b \ 0\ div_mult_self1 [of b a c] by (simp add: ac_simps) + +lemma div_mult_self3: + "(c * b + a) div b = c + a div b" + using \b \ 0\ div_mult_self1 [of b a c] by (simp add: ac_simps) + +lemma div_mult_self4: + "(b * c + a) div b = c + a div b" + using \b \ 0\ div_mult_self1 [of b a c] by (simp add: ac_simps) + +lemma div_add_self1: + "(b + a) div b = a div b + 1" + using \b \ 0\ div_mult_self1 [of b a 1] by (simp add: ac_simps) + +lemma div_add_self2: + "(a + b) div b = a div b + 1" + using \b \ 0\ div_add_self1 [of a] by (simp add: ac_simps) + +end + +lemma div_mult_mult2: + "(a * c) div (b * c) = a div b" if "c \ 0" + using that div_mult_mult1 [of c a b] by (simp add: ac_simps) + +lemma div_mult_mult1_if [simp]: + "(c * a) div (c * b) = (if c = 0 then 0 else a div b)" + by (simp add: div_mult_mult1) + +lemma div_mult_mult2_if [simp]: + "(a * c) div (b * c) = (if c = 0 then 0 else a div b)" + using div_mult_mult1_if [of c a b] by (simp add: ac_simps) + +end + +class idom_divide_cancel = idom_divide + semidom_divide_cancel +begin + +lemma div_minus_minus [simp]: "(- a) div (- b) = a div b" + using div_mult_mult1 [of "- 1" a b] by simp + +lemma div_minus_right: "a div (- b) = (- a) div b" + using div_minus_minus [of "- a" b] by simp + +lemma div_minus1_right [simp]: "a div (- 1) = - a" + using div_minus_right [of a 1] by simp + +end + + +subsection \Basic notions following from divisibility\ + class algebraic_semidom = semidom_divide begin @@ -1699,70 +1761,7 @@ end -text \Integral (semi)domains with cancellation rules\ - -class semidom_divide_cancel = semidom_divide + - assumes div_mult_self1: "b \ 0 \ (a + c * b) div b = c + a div b" - and div_mult_mult1: "c \ 0 \ (c * a) div (c * b) = a div b" -begin - -context - fixes b - assumes "b \ 0" -begin - -lemma div_mult_self2: - "(a + b * c) div b = c + a div b" - using \b \ 0\ div_mult_self1 [of b a c] by (simp add: ac_simps) - -lemma div_mult_self3: - "(c * b + a) div b = c + a div b" - using \b \ 0\ div_mult_self1 [of b a c] by (simp add: ac_simps) - -lemma div_mult_self4: - "(b * c + a) div b = c + a div b" - using \b \ 0\ div_mult_self1 [of b a c] by (simp add: ac_simps) - -lemma div_add_self1: - "(b + a) div b = a div b + 1" - using \b \ 0\ div_mult_self1 [of b a 1] by (simp add: ac_simps) - -lemma div_add_self2: - "(a + b) div b = a div b + 1" - using \b \ 0\ div_add_self1 [of a] by (simp add: ac_simps) - -end - -lemma div_mult_mult2: - "(a * c) div (b * c) = a div b" if "c \ 0" - using that div_mult_mult1 [of c a b] by (simp add: ac_simps) - -lemma div_mult_mult1_if [simp]: - "(c * a) div (c * b) = (if c = 0 then 0 else a div b)" - by (simp add: div_mult_mult1) - -lemma div_mult_mult2_if [simp]: - "(a * c) div (b * c) = (if c = 0 then 0 else a div b)" - using div_mult_mult1_if [of c a b] by (simp add: ac_simps) - -end - -class idom_divide_cancel = idom_divide + semidom_divide_cancel -begin - -lemma div_minus_minus [simp]: "(- a) div (- b) = a div b" - using div_mult_mult1 [of "- 1" a b] by simp - -lemma div_minus_right: "a div (- b) = (- a) div b" - using div_minus_minus [of "- a" b] by simp - -lemma div_minus1_right [simp]: "a div (- 1) = - a" - using div_minus_right [of a 1] by simp - -end - - -text \Quotient and remainder in integral domains\ +subsection \Quotient and remainder in integral domains\ class semidom_modulo = algebraic_semidom + semiring_modulo begin @@ -1833,7 +1832,19 @@ end -text \Interlude: basic tool support for algebraic and arithmetic calculations\ +class idom_modulo = idom + semidom_modulo +begin + +subclass idom_divide .. + +lemma div_diff [simp]: + "c dvd a \ c dvd b \ (a - b) div c = a div c - b div c" + using div_add [of _ _ "- b"] by (simp add: dvd_neg_div) + +end + + +subsection \Interlude: basic tool support for algebraic and arithmetic calculations\ named_theorems arith "arith facts -- only ground formulas" ML_file \Tools/arith_data.ML\ @@ -1859,17 +1870,18 @@ simproc_setup cancel_div_mod_int ("(a::'a::semidom_modulo) + b") = \K Cancel_Div_Mod_Ring.proc\ -class idom_modulo = idom + semidom_modulo -begin - -subclass idom_divide .. - -lemma div_diff [simp]: - "c dvd a \ c dvd b \ (a - b) div c = a div c - b div c" - using div_add [of _ _ "- b"] by (simp add: dvd_neg_div) - -end - + +subsection \Ordered semirings and rings\ + +text \ + The theory of partially ordered rings is taken from the books: + \<^item> \<^emph>\Lattice Theory\ by Garret Birkhoff, American Mathematical Society, 1979 + \<^item> \<^emph>\Partially Ordered Algebraic Systems\, Pergamon Press, 1963 + + Most of the used notions can also be looked up in + \<^item> \<^url>\http://www.mathworld.com\ by Eric Weisstein et. al. + \<^item> \<^emph>\Algebra I\ by van der Waerden, Springer +\ class ordered_semiring = semiring + ordered_comm_monoid_add + assumes mult_left_mono: "a \ b \ 0 \ c \ c * a \ c * b" @@ -2669,6 +2681,7 @@ end + subsection \Dioids\ text \