src/HOL/Algebra/abstract/Ring2.thy
author wenzelm
Wed, 16 Jul 2008 11:20:25 +0200
changeset 27617 dee36037a832
parent 27542 9bf0a22f8bcd
child 27651 16a26996c30e
permissions -rw-r--r--
added Isar.command, Isar.insert, Isar.remove (editor model);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
     1
(*
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
     2
  Title:     The algebraic hierarchy of rings as axiomatic classes
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
     3
  Id:        $Id$
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
     4
  Author:    Clemens Ballarin, started 9 December 1996
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
     5
  Copyright: Clemens Ballarin
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
     6
*)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
     7
27542
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
     8
header {* The algebraic hierarchy of rings as type classes *}
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
     9
27542
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
    10
theory Ring2
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
    11
imports Main
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
    12
begin
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    13
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    14
subsection {* Ring axioms *}
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    15
27542
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
    16
class ring = zero + one + plus + minus + uminus + times + inverse + power + Divides.dvd +
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
    17
  assumes a_assoc:      "(a + b) + c = a + (b + c)"
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
    18
  and l_zero:           "0 + a = a"
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
    19
  and l_neg:            "(-a) + a = 0"
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
    20
  and a_comm:           "a + b = b + a"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    21
27542
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
    22
  assumes m_assoc:      "(a * b) * c = a * (b * c)"
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
    23
  and l_one:            "1 * a = a"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    24
27542
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
    25
  assumes l_distr:      "(a + b) * c = a * c + b * c"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    26
27542
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
    27
  assumes m_comm:       "a * b = b * a"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    28
27542
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
    29
  assumes minus_def:    "a - b = a + (-b)"
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
    30
  and inverse_def:      "inverse a = (if a dvd 1 then THE x. a*x = 1 else 0)"
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
    31
  and divide_def:       "a / b = a * inverse b"
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
    32
  and power_0 [simp]:   "a ^ 0 = 1"
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
    33
  and power_Suc [simp]: "a ^ Suc n = a ^ n * a"
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
    34
begin
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    35
27542
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
    36
definition assoc :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl "assoc" 50) where
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
    37
  assoc_def: "a assoc b \<longleftrightarrow> a dvd b & b dvd a"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    38
27542
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
    39
definition irred :: "'a \<Rightarrow> bool" where
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
    40
  irred_def: "irred a \<longleftrightarrow> a ~= 0 & ~ a dvd 1
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    41
                          & (ALL d. d dvd a --> d dvd 1 | a dvd d)"
27542
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
    42
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
    43
definition prime :: "'a \<Rightarrow> bool" where
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
    44
  prime_def: "prime p \<longleftrightarrow> p ~= 0 & ~ p dvd 1
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    45
                          & (ALL a b. p dvd (a*b) --> p dvd a | p dvd b)"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    46
27542
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
    47
end
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
    48
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
    49
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    50
subsection {* Integral domains *}
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    51
27542
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
    52
class "domain" = ring +
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
    53
  assumes one_not_zero: "1 ~= 0"
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
    54
  and integral: "a * b = 0 ==> a = 0 | b = 0"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    55
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    56
subsection {* Factorial domains *}
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    57
27542
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
    58
class factorial = "domain" +
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    59
(*
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    60
  Proper definition using divisor chain condition currently not supported.
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    61
  factorial_divisor:    "wf {(a, b). a dvd b & ~ (b dvd a)}"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    62
*)
27542
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
    63
  assumes factorial_divisor: "True"
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
    64
  and factorial_prime: "irred a ==> prime a"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    65
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    66
subsection {* Euclidean domains *}
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    67
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    68
(*
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    69
axclass
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    70
  euclidean < "domain"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    71
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    72
  euclidean_ax:  "b ~= 0 ==> Ex (% (q, r, e_size::('a::ringS)=>nat).
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    73
                   a = b * q + r & e_size r < e_size b)"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    74
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    75
  Nothing has been proved about Euclidean domains, yet.
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    76
  Design question:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    77
    Fix quo, rem and e_size as constants that are axiomatised with
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    78
    euclidean_ax?
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    79
    - advantage: more pragmatic and easier to use
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    80
    - disadvantage: for every type, one definition of quo and rem will
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    81
        be fixed, users may want to use differing ones;
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    82
        also, it seems not possible to prove that fields are euclidean
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    83
        domains, because that would require generic (type-independent)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    84
        definitions of quo and rem.
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    85
*)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    86
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    87
subsection {* Fields *}
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    88
27542
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
    89
class field = ring +
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
    90
  assumes field_one_not_zero: "1 ~= 0"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    91
                (* Avoid a common superclass as the first thing we will
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    92
                   prove about fields is that they are domains. *)
27542
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
    93
  and field_ax: "a ~= 0 ==> a dvd 1"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    94
21423
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
    95
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    96
section {* Basic facts *}
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    97
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    98
subsection {* Normaliser for rings *}
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    99
21423
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   100
(* derived rewrite rules *)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   101
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   102
lemma a_lcomm: "(a::'a::ring)+(b+c) = b+(a+c)"
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   103
  apply (rule a_comm [THEN trans])
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   104
  apply (rule a_assoc [THEN trans])
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   105
  apply (rule a_comm [THEN arg_cong])
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   106
  done
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   107
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   108
lemma r_zero: "(a::'a::ring) + 0 = a"
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   109
  apply (rule a_comm [THEN trans])
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   110
  apply (rule l_zero)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   111
  done
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   112
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   113
lemma r_neg: "(a::'a::ring) + (-a) = 0"
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   114
  apply (rule a_comm [THEN trans])
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   115
  apply (rule l_neg)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   116
  done
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   117
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   118
lemma r_neg2: "(a::'a::ring) + (-a + b) = b"
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   119
  apply (rule a_assoc [symmetric, THEN trans])
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   120
  apply (simp add: r_neg l_zero)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   121
  done
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   122
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   123
lemma r_neg1: "-(a::'a::ring) + (a + b) = b"
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   124
  apply (rule a_assoc [symmetric, THEN trans])
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   125
  apply (simp add: l_neg l_zero)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   126
  done
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   127
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   128
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   129
(* auxiliary *)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   130
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   131
lemma a_lcancel: "!! a::'a::ring. a + b = a + c ==> b = c"
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   132
  apply (rule box_equals)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   133
  prefer 2
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   134
  apply (rule l_zero)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   135
  prefer 2
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   136
  apply (rule l_zero)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   137
  apply (rule_tac a1 = a in l_neg [THEN subst])
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   138
  apply (simp add: a_assoc)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   139
  done
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   140
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   141
lemma minus_add: "-((a::'a::ring) + b) = (-a) + (-b)"
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   142
  apply (rule_tac a = "a + b" in a_lcancel)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   143
  apply (simp add: r_neg l_neg l_zero a_assoc a_comm a_lcomm)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   144
  done
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   145
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   146
lemma minus_minus: "-(-(a::'a::ring)) = a"
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   147
  apply (rule a_lcancel)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   148
  apply (rule r_neg [THEN trans])
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   149
  apply (rule l_neg [symmetric])
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   150
  done
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   151
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   152
lemma minus0: "- 0 = (0::'a::ring)"
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   153
  apply (rule a_lcancel)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   154
  apply (rule r_neg [THEN trans])
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   155
  apply (rule l_zero [symmetric])
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   156
  done
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   157
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   158
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   159
(* derived rules for multiplication *)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   160
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   161
lemma m_lcomm: "(a::'a::ring)*(b*c) = b*(a*c)"
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   162
  apply (rule m_comm [THEN trans])
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   163
  apply (rule m_assoc [THEN trans])
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   164
  apply (rule m_comm [THEN arg_cong])
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   165
  done
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   166
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   167
lemma r_one: "(a::'a::ring) * 1 = a"
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   168
  apply (rule m_comm [THEN trans])
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   169
  apply (rule l_one)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   170
  done
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   171
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   172
lemma r_distr: "(a::'a::ring) * (b + c) = a * b + a * c"
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   173
  apply (rule m_comm [THEN trans])
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   174
  apply (rule l_distr [THEN trans])
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   175
  apply (simp add: m_comm)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   176
  done
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   177
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   178
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   179
(* the following proof is from Jacobson, Basic Algebra I, pp. 88-89 *)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   180
lemma l_null: "0 * (a::'a::ring) = 0"
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   181
  apply (rule a_lcancel)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   182
  apply (rule l_distr [symmetric, THEN trans])
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   183
  apply (simp add: r_zero)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   184
  done
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   185
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   186
lemma r_null: "(a::'a::ring) * 0 = 0"
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   187
  apply (rule m_comm [THEN trans])
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   188
  apply (rule l_null)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   189
  done
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   190
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   191
lemma l_minus: "(-(a::'a::ring)) * b = - (a * b)"
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   192
  apply (rule a_lcancel)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   193
  apply (rule r_neg [symmetric, THEN [2] trans])
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   194
  apply (rule l_distr [symmetric, THEN trans])
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   195
  apply (simp add: l_null r_neg)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   196
  done
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   197
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   198
lemma r_minus: "(a::'a::ring) * (-b) = - (a * b)"
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   199
  apply (rule a_lcancel)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   200
  apply (rule r_neg [symmetric, THEN [2] trans])
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   201
  apply (rule r_distr [symmetric, THEN trans])
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   202
  apply (simp add: r_null r_neg)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   203
  done
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   204
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   205
(*** Term order for commutative rings ***)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   206
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   207
ML {*
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   208
fun ring_ord (Const (a, _)) =
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   209
    find_index (fn a' => a = a')
22997
d4f3b015b50b canonical prefixing of class constants
haftmann
parents: 22808
diff changeset
   210
      [@{const_name HOL.zero}, @{const_name HOL.plus}, @{const_name HOL.uminus},
d4f3b015b50b canonical prefixing of class constants
haftmann
parents: 22808
diff changeset
   211
        @{const_name HOL.minus}, @{const_name HOL.one}, @{const_name HOL.times}]
21423
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   212
  | ring_ord _ = ~1;
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   213
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   214
fun termless_ring (a, b) = (Term.term_lpo ring_ord (a, b) = LESS);
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   215
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   216
val ring_ss = HOL_basic_ss settermless termless_ring addsimps
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   217
  [thm "a_assoc", thm "l_zero", thm "l_neg", thm "a_comm", thm "m_assoc",
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   218
   thm "l_one", thm "l_distr", thm "m_comm", thm "minus_def",
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   219
   thm "r_zero", thm "r_neg", thm "r_neg2", thm "r_neg1", thm "minus_add",
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   220
   thm "minus_minus", thm "minus0", thm "a_lcomm", thm "m_lcomm", (*thm "r_one",*)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   221
   thm "r_distr", thm "l_null", thm "r_null", thm "l_minus", thm "r_minus"];
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   222
*}   (* Note: r_one is not necessary in ring_ss *)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   223
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   224
method_setup ring =
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21423
diff changeset
   225
  {* Method.no_args (Method.SIMPLE_METHOD' (full_simp_tac ring_ss)) *}
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   226
  {* computes distributive normal form in rings *}
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   227
21423
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   228
lemmas ring_simps =
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   229
  l_zero r_zero l_neg r_neg minus_minus minus0
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   230
  l_one r_one l_null r_null l_minus r_minus
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   231
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   232
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   233
subsection {* Rings and the summation operator *}
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   234
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   235
(* Basic facts --- move to HOL!!! *)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   236
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   237
(* needed because natsum_cong (below) disables atMost_0 *)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   238
lemma natsum_0 [simp]: "setsum f {..(0::nat)} = (f 0::'a::comm_monoid_add)"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   239
by simp
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   240
(*
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   241
lemma natsum_Suc [simp]:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   242
  "setsum f {..Suc n} = (f (Suc n) + setsum f {..n}::'a::comm_monoid_add)"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   243
by (simp add: atMost_Suc)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   244
*)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   245
lemma natsum_Suc2:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   246
  "setsum f {..Suc n} = (f 0::'a::comm_monoid_add) + (setsum (%i. f (Suc i)) {..n})"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   247
proof (induct n)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   248
  case 0 show ?case by simp
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   249
next
22384
33a46e6c7f04 prefix of class interpretation not mandatory any longer
haftmann
parents: 21588
diff changeset
   250
  case Suc thus ?case by (simp add: add_assoc) 
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   251
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   252
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   253
lemma natsum_cong [cong]:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   254
  "!!k. [| j = k; !!i::nat. i <= k ==> f i = (g i::'a::comm_monoid_add) |] ==>
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   255
        setsum f {..j} = setsum g {..k}"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   256
by (induct j) auto
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   257
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   258
lemma natsum_zero [simp]: "setsum (%i. 0) {..n::nat} = (0::'a::comm_monoid_add)"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   259
by (induct n) simp_all
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   260
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   261
lemma natsum_add [simp]:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   262
  "!!f::nat=>'a::comm_monoid_add.
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   263
   setsum (%i. f i + g i) {..n::nat} = setsum f {..n} + setsum g {..n}"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   264
by (induct n) (simp_all add: add_ac)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   265
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   266
(* Facts specific to rings *)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   267
27542
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
   268
subclass (in ring) comm_monoid_add
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   269
proof
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   270
  fix x y z
27542
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
   271
  show "x + y = y + x" by (rule a_comm)
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
   272
  show "(x + y) + z = x + (y + z)" by (rule a_assoc)
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
   273
  show "0 + x = x" by (rule l_zero)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   274
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   275
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   276
ML {*
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   277
  local
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   278
    val lhss = 
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   279
        ["t + u::'a::ring",
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   280
	 "t - u::'a::ring",
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   281
	 "t * u::'a::ring",
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   282
	 "- t::'a::ring"];
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   283
    fun proc ss t = 
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   284
      let val rew = Goal.prove (Simplifier.the_context ss) [] []
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   285
            (HOLogic.mk_Trueprop
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   286
              (HOLogic.mk_eq (t, Var (("x", Term.maxidx_of_term t + 1), fastype_of t))))
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   287
                (fn _ => simp_tac (Simplifier.inherit_context ss ring_ss) 1)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   288
            |> mk_meta_eq;
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   289
          val (t', u) = Logic.dest_equals (Thm.prop_of rew);
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   290
      in if t' aconv u 
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   291
        then NONE
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   292
        else SOME rew 
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   293
    end;
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   294
  in
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   295
    val ring_simproc = Simplifier.simproc (the_context ()) "ring" lhss (K proc);
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   296
  end;
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   297
*}
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   298
26480
544cef16045b replaced 'ML_setup' by 'ML';
wenzelm
parents: 26342
diff changeset
   299
ML {* Addsimprocs [ring_simproc] *}
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   300
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   301
lemma natsum_ldistr:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   302
  "!!a::'a::ring. setsum f {..n::nat} * a = setsum (%i. f i * a) {..n}"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   303
by (induct n) simp_all
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   304
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   305
lemma natsum_rdistr:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   306
  "!!a::'a::ring. a * setsum f {..n::nat} = setsum (%i. a * f i) {..n}"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   307
by (induct n) simp_all
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   308
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   309
subsection {* Integral Domains *}
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   310
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   311
declare one_not_zero [simp]
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   312
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   313
lemma zero_not_one [simp]:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   314
  "0 ~= (1::'a::domain)" 
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   315
by (rule not_sym) simp
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   316
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   317
lemma integral_iff: (* not by default a simp rule! *)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   318
  "(a * b = (0::'a::domain)) = (a = 0 | b = 0)"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   319
proof
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   320
  assume "a * b = 0" then show "a = 0 | b = 0" by (simp add: integral)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   321
next
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   322
  assume "a = 0 | b = 0" then show "a * b = 0" by auto
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   323
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   324
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   325
(*
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   326
lemma "(a::'a::ring) - (a - b) = b" apply simp
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   327
 simproc seems to fail on this example (fixed with new term order)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   328
*)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   329
(*
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   330
lemma bug: "(b::'a::ring) - (b - a) = a" by simp
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   331
   simproc for rings cannot prove "(a::'a::ring) - (a - b) = b" 
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   332
*)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   333
lemma m_lcancel:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   334
  assumes prem: "(a::'a::domain) ~= 0" shows conc: "(a * b = a * c) = (b = c)"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   335
proof
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   336
  assume eq: "a * b = a * c"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   337
  then have "a * (b - c) = 0" by simp
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   338
  then have "a = 0 | (b - c) = 0" by (simp only: integral_iff)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   339
  with prem have "b - c = 0" by auto 
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   340
  then have "b = b - (b - c)" by simp 
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   341
  also have "b - (b - c) = c" by simp
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   342
  finally show "b = c" .
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   343
next
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   344
  assume "b = c" then show "a * b = a * c" by simp
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   345
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   346
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   347
lemma m_rcancel:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   348
  "(a::'a::domain) ~= 0 ==> (b * a = c * a) = (b = c)"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   349
by (simp add: m_lcancel)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   350
27542
9bf0a22f8bcd class instead of axclass
haftmann
parents: 26480
diff changeset
   351
declare power_Suc [simp]
21416
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   352
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   353
lemma power_one [simp]:
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   354
  "1 ^ n = (1::'a::ring)" by (induct n) simp_all
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   355
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   356
lemma power_zero [simp]:
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   357
  "n \<noteq> 0 \<Longrightarrow> 0 ^ n = (0::'a::ring)" by (induct n) simp_all
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   358
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   359
lemma power_mult [simp]:
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   360
  "(a::'a::ring) ^ m * a ^ n = a ^ (m + n)"
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   361
  by (induct m) simp_all
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   362
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   363
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   364
section "Divisibility"
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   365
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   366
lemma dvd_zero_right [simp]:
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   367
  "(a::'a::ring) dvd 0"
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   368
proof
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   369
  show "0 = a * 0" by simp
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   370
qed
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   371
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   372
lemma dvd_zero_left:
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   373
  "0 dvd (a::'a::ring) \<Longrightarrow> a = 0" unfolding dvd_def by simp
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   374
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   375
lemma dvd_refl_ring [simp]:
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   376
  "(a::'a::ring) dvd a"
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   377
proof
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   378
  show "a = a * 1" by simp
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   379
qed
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   380
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   381
lemma dvd_trans_ring:
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   382
  fixes a b c :: "'a::ring"
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   383
  assumes a_dvd_b: "a dvd b"
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   384
  and b_dvd_c: "b dvd c"
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   385
  shows "a dvd c"
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   386
proof -
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   387
  from a_dvd_b obtain l where "b = a * l" using dvd_def by blast
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   388
  moreover from b_dvd_c obtain j where "c = b * j" using dvd_def by blast
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   389
  ultimately have "c = a * (l * j)" by simp
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   390
  then have "\<exists>k. c = a * k" ..
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   391
  then show ?thesis using dvd_def by blast
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   392
qed
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   393
21423
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   394
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   395
lemma unit_mult: 
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   396
  "!!a::'a::ring. [| a dvd 1; b dvd 1 |] ==> a * b dvd 1"
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   397
  apply (unfold dvd_def)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   398
  apply clarify
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   399
  apply (rule_tac x = "k * ka" in exI)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   400
  apply simp
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   401
  done
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   402
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   403
lemma unit_power: "!!a::'a::ring. a dvd 1 ==> a^n dvd 1"
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   404
  apply (induct_tac n)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   405
   apply simp
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   406
  apply (simp add: unit_mult)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   407
  done
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   408
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   409
lemma dvd_add_right [simp]:
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   410
  "!! a::'a::ring. [| a dvd b; a dvd c |] ==> a dvd b + c"
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   411
  apply (unfold dvd_def)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   412
  apply clarify
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   413
  apply (rule_tac x = "k + ka" in exI)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   414
  apply (simp add: r_distr)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   415
  done
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   416
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   417
lemma dvd_uminus_right [simp]:
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   418
  "!! a::'a::ring. a dvd b ==> a dvd -b"
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   419
  apply (unfold dvd_def)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   420
  apply clarify
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   421
  apply (rule_tac x = "-k" in exI)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   422
  apply (simp add: r_minus)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   423
  done
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   424
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   425
lemma dvd_l_mult_right [simp]:
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   426
  "!! a::'a::ring. a dvd b ==> a dvd c*b"
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   427
  apply (unfold dvd_def)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   428
  apply clarify
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   429
  apply (rule_tac x = "c * k" in exI)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   430
  apply simp
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   431
  done
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   432
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   433
lemma dvd_r_mult_right [simp]:
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   434
  "!! a::'a::ring. a dvd b ==> a dvd b*c"
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   435
  apply (unfold dvd_def)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   436
  apply clarify
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   437
  apply (rule_tac x = "k * c" in exI)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   438
  apply simp
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   439
  done
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   440
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   441
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   442
(* Inverse of multiplication *)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   443
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   444
section "inverse"
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   445
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   446
lemma inverse_unique: "!! a::'a::ring. [| a * x = 1; a * y = 1 |] ==> x = y"
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   447
  apply (rule_tac a = "(a*y) * x" and b = "y * (a*x)" in box_equals)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   448
    apply (simp (no_asm))
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   449
  apply auto
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   450
  done
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   451
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   452
lemma r_inverse_ring: "!! a::'a::ring. a dvd 1 ==> a * inverse a = 1"
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   453
  apply (unfold inverse_def dvd_def)
26342
0f65fa163304 more antiquotations;
wenzelm
parents: 25762
diff changeset
   454
  apply (tactic {* asm_full_simp_tac (@{simpset} delsimprocs [ring_simproc]) 1 *})
21423
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   455
  apply clarify
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   456
  apply (rule theI)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   457
   apply assumption
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   458
  apply (rule inverse_unique)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   459
   apply assumption
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   460
  apply assumption
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   461
  done
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   462
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   463
lemma l_inverse_ring: "!! a::'a::ring. a dvd 1 ==> inverse a * a = 1"
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   464
  by (simp add: r_inverse_ring)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   465
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   466
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   467
(* Fields *)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   468
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   469
section "Fields"
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   470
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   471
lemma field_unit [simp]: "!! a::'a::field. (a dvd 1) = (a ~= 0)"
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   472
  by (auto dest: field_ax dvd_zero_left simp add: field_one_not_zero)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   473
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   474
lemma r_inverse [simp]: "!! a::'a::field. a ~= 0 ==> a * inverse a = 1"
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   475
  by (simp add: r_inverse_ring)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   476
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   477
lemma l_inverse [simp]: "!! a::'a::field. a ~= 0 ==> inverse a * a= 1"
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   478
  by (simp add: l_inverse_ring)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   479
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   480
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   481
(* fields are integral domains *)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   482
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   483
lemma field_integral: "!! a::'a::field. a * b = 0 ==> a = 0 | b = 0"
23894
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 22997
diff changeset
   484
  apply (tactic "step_tac @{claset} 1")
21423
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   485
  apply (rule_tac a = " (a*b) * inverse b" in box_equals)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   486
    apply (rule_tac [3] refl)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   487
   prefer 2
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   488
   apply (simp (no_asm))
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   489
   apply auto
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   490
  done
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   491
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   492
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   493
(* fields are factorial domains *)
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   494
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   495
lemma field_fact_prime: "!! a::'a::field. irred a ==> prime a"
6cdd0589aa73 HOL-Algebra: converted legacy ML scripts;
wenzelm
parents: 21416
diff changeset
   496
  unfolding prime_def irred_def by (blast intro: field_ax)
21416
f23e4e75dfd3 dvd_def now with object equality
haftmann
parents: 20318
diff changeset
   497
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   498
end