src/HOL/Algebra/abstract/Ring.ML
author ballarin
Thu, 28 Nov 2002 10:50:42 +0100
changeset 13735 7de9342aca7a
parent 13550 5a176b8dda84
child 13783 3294f727e20d
permissions -rw-r--r--
HOL-Algebra partially ported to Isar.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
     1
(*
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
     2
    Abstract class ring (commutative, with 1)
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
     3
    $Id$
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
     4
    Author: Clemens Ballarin, started 9 December 1996
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
     5
*)
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
     6
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
     7
(*
11778
37efbe093d3c ring includes plus_ac0;
wenzelm
parents: 11093
diff changeset
     8
val a_assoc = thm "plus_ac0.assoc";
37efbe093d3c ring includes plus_ac0;
wenzelm
parents: 11093
diff changeset
     9
val l_zero = thm "plus_ac0.zero";
37efbe093d3c ring includes plus_ac0;
wenzelm
parents: 11093
diff changeset
    10
val a_comm = thm "plus_ac0.commute";
37efbe093d3c ring includes plus_ac0;
wenzelm
parents: 11093
diff changeset
    11
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    12
section "Rings";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    13
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    14
fun make_left_commute assoc commute s =
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    15
  [rtac (commute RS trans) 1, rtac (assoc RS trans) 1,
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    16
   rtac (commute RS arg_cong) 1];
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    17
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    18
(* addition *)
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    19
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    20
qed_goal "a_lcomm" Ring.thy "!!a::'a::ring. a+(b+c) = b+(a+c)"
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    21
  (make_left_commute a_assoc a_comm);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    22
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    23
val a_ac = [a_assoc, a_comm, a_lcomm];
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    24
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10789
diff changeset
    25
Goal "!!a::'a::ring. a + 0 = a";
9390
e6b96d953965 renamed // to / (which is what we want anyway) to avoid clash with the new
paulson
parents: 8707
diff changeset
    26
by (rtac (a_comm RS trans) 1);
e6b96d953965 renamed // to / (which is what we want anyway) to avoid clash with the new
paulson
parents: 8707
diff changeset
    27
by (rtac l_zero 1);
e6b96d953965 renamed // to / (which is what we want anyway) to avoid clash with the new
paulson
parents: 8707
diff changeset
    28
qed "r_zero";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    29
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10789
diff changeset
    30
Goal "!!a::'a::ring. a + (-a) = 0";
9390
e6b96d953965 renamed // to / (which is what we want anyway) to avoid clash with the new
paulson
parents: 8707
diff changeset
    31
by (rtac (a_comm RS trans) 1);
e6b96d953965 renamed // to / (which is what we want anyway) to avoid clash with the new
paulson
parents: 8707
diff changeset
    32
by (rtac l_neg 1);
e6b96d953965 renamed // to / (which is what we want anyway) to avoid clash with the new
paulson
parents: 8707
diff changeset
    33
qed "r_neg";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    34
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    35
Goal "!! a::'a::ring. a + b = a + c ==> b = c";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    36
by (rtac box_equals 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    37
by (rtac l_zero 2);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    38
by (rtac l_zero 2);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    39
by (res_inst_tac [("a1", "a")] (l_neg RS subst) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    40
by (asm_simp_tac (simpset() addsimps [a_assoc]) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    41
qed "a_lcancel";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    42
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    43
Goal "!! a::'a::ring. b + a = c + a ==> b = c";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    44
by (rtac a_lcancel 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    45
by (asm_simp_tac (simpset() addsimps a_ac) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    46
qed "a_rcancel";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    47
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    48
Goal "!! a::'a::ring. (a + b = a + c) = (b = c)";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    49
by (auto_tac (claset() addSDs [a_lcancel], simpset()));
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    50
qed "a_lcancel_eq";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    51
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    52
Goal "!! a::'a::ring. (b + a = c + a) = (b = c)";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    53
by (simp_tac (simpset() addsimps [a_lcancel_eq, a_comm]) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    54
qed "a_rcancel_eq";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    55
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    56
Addsimps [a_lcancel_eq, a_rcancel_eq];
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    57
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    58
Goal "!!a::'a::ring. -(-a) = a";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    59
by (rtac a_lcancel 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    60
by (rtac (r_neg RS trans) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    61
by (rtac (l_neg RS sym) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    62
qed "minus_minus";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    63
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10789
diff changeset
    64
Goal "- 0 = (0::'a::ring)";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    65
by (rtac a_lcancel 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    66
by (rtac (r_neg RS trans) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    67
by (rtac (l_zero RS sym) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    68
qed "minus0";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    69
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    70
Goal "!!a::'a::ring. -(a + b) = (-a) + (-b)";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    71
by (res_inst_tac [("a", "a+b")] a_lcancel 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    72
by (simp_tac (simpset() addsimps ([r_neg, l_neg, l_zero]@a_ac)) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    73
qed "minus_add";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    74
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    75
(* multiplication *)
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    76
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    77
qed_goal "m_lcomm" Ring.thy "!!a::'a::ring. a*(b*c) = b*(a*c)"
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    78
  (make_left_commute m_assoc m_comm);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    79
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    80
val m_ac = [m_assoc, m_comm, m_lcomm];
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    81
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
    82
Goal "!!a::'a::ring. a * 1 = a";
9390
e6b96d953965 renamed // to / (which is what we want anyway) to avoid clash with the new
paulson
parents: 8707
diff changeset
    83
by (rtac (m_comm RS trans) 1);
e6b96d953965 renamed // to / (which is what we want anyway) to avoid clash with the new
paulson
parents: 8707
diff changeset
    84
by (rtac l_one 1);
e6b96d953965 renamed // to / (which is what we want anyway) to avoid clash with the new
paulson
parents: 8707
diff changeset
    85
qed "r_one";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    86
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    87
(* distributive and derived *)
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    88
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    89
Goal "!!a::'a::ring. a * (b + c) = a * b + a * c";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    90
by (rtac (m_comm RS trans) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    91
by (rtac (l_distr RS trans) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    92
by (simp_tac (simpset() addsimps [m_comm]) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    93
qed "r_distr";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    94
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    95
val m_distr = m_ac @ [l_distr, r_distr];
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    96
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    97
(* the following two proofs can be found in
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    98
   Jacobson, Basic Algebra I, pp. 88-89 *)
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    99
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10789
diff changeset
   100
Goal "!!a::'a::ring. 0 * a = 0";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   101
by (rtac a_lcancel 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   102
by (rtac (l_distr RS sym RS trans) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   103
by (simp_tac (simpset() addsimps [r_zero]) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   104
qed "l_null";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   105
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10789
diff changeset
   106
Goal "!!a::'a::ring. a * 0 = 0";
9390
e6b96d953965 renamed // to / (which is what we want anyway) to avoid clash with the new
paulson
parents: 8707
diff changeset
   107
by (rtac (m_comm RS trans) 1);
e6b96d953965 renamed // to / (which is what we want anyway) to avoid clash with the new
paulson
parents: 8707
diff changeset
   108
by (rtac l_null 1);
e6b96d953965 renamed // to / (which is what we want anyway) to avoid clash with the new
paulson
parents: 8707
diff changeset
   109
qed "r_null";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   110
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   111
Goal "!!a::'a::ring. (-a) * b = - (a * b)";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   112
by (rtac a_lcancel 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   113
by (rtac (r_neg RS sym RSN (2, trans)) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   114
by (rtac (l_distr RS sym RS trans) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   115
by (simp_tac (simpset() addsimps [l_null, r_neg]) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   116
qed "l_minus";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   117
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   118
Goal "!!a::'a::ring. a * (-b) = - (a * b)";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   119
by (rtac a_lcancel 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   120
by (rtac (r_neg RS sym RSN (2, trans)) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   121
by (rtac (r_distr RS sym RS trans) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   122
by (simp_tac (simpset() addsimps [r_null, r_neg]) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   123
qed "r_minus";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   124
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   125
val m_minus = [l_minus, r_minus];
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   126
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   127
Addsimps [l_zero, r_zero, l_neg, r_neg, minus_minus, minus0, 
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10789
diff changeset
   128
	  l_one, r_one, l_null, r_null];
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   129
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   130
(* further rules *)
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   131
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10789
diff changeset
   132
Goal "!!a::'a::ring. -a = 0 ==> a = 0";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   133
by (res_inst_tac [("t", "a")] (minus_minus RS subst) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   134
by (Asm_simp_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   135
qed "uminus_monom";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   136
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10789
diff changeset
   137
Goal "!!a::'a::ring. a ~= 0 ==> -a ~= 0";
10198
paulson
parents: 9970
diff changeset
   138
by (blast_tac (claset() addIs [uminus_monom]) 1); 
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   139
qed "uminus_monom_neq";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   140
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10789
diff changeset
   141
Goal "!!a::'a::ring. a * b ~= 0 ==> a ~= 0";
10198
paulson
parents: 9970
diff changeset
   142
by Auto_tac;  
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   143
qed "l_nullD";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   144
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10789
diff changeset
   145
Goal "!!a::'a::ring. a * b ~= 0 ==> b ~= 0";
10198
paulson
parents: 9970
diff changeset
   146
by Auto_tac;  
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   147
qed "r_nullD";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   148
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10789
diff changeset
   149
(* reflection between a = b and a -- b = 0 *)
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   150
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10789
diff changeset
   151
Goal "!!a::'a::ring. a = b ==> a + (-b) = 0";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   152
by (Asm_simp_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   153
qed "eq_imp_diff_zero";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   154
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10789
diff changeset
   155
Goal "!!a::'a::ring. a + (-b) = 0 ==> a = b";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   156
by (res_inst_tac [("a", "-b")] a_rcancel 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   157
by (Asm_simp_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   158
qed "diff_zero_imp_eq";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   159
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   160
(* this could be a rewrite rule, but won't terminate
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   161
   ==> make it a simproc?
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10789
diff changeset
   162
Goal "!!a::'a::ring. (a = b) = (a -- b = 0)";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   163
*)
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   164
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   165
*)
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   166
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   167
(* Power *)
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   168
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   169
Goal "!!a::'a::ring. a ^ 0 = 1";
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   170
by (simp_tac (simpset() addsimps [power_def]) 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   171
qed "power_0";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   172
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   173
Goal "!!a::'a::ring. a ^ Suc n = a ^ n * a";
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   174
by (simp_tac (simpset() addsimps [power_def]) 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   175
qed "power_Suc";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   176
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   177
Addsimps [power_0, power_Suc];
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   178
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   179
Goal "1 ^ n = (1::'a::ring)";
8707
paulson
parents: 7998
diff changeset
   180
by (induct_tac "n" 1);
paulson
parents: 7998
diff changeset
   181
by Auto_tac;
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   182
qed "power_one";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   183
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10789
diff changeset
   184
Goal "!!n. n ~= 0 ==> 0 ^ n = (0::'a::ring)";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   185
by (etac rev_mp 1);
8707
paulson
parents: 7998
diff changeset
   186
by (induct_tac "n" 1);
paulson
parents: 7998
diff changeset
   187
by Auto_tac;
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   188
qed "power_zero";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   189
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   190
Addsimps [power_zero, power_one];
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   191
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   192
Goal "!! a::'a::ring. a ^ m * a ^ n = a ^ (m + n)";
8707
paulson
parents: 7998
diff changeset
   193
by (induct_tac "m" 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   194
by (Simp_tac 1);
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   195
by (Asm_simp_tac 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   196
qed "power_mult";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   197
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   198
(* Divisibility *)
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   199
section "Divisibility";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   200
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10789
diff changeset
   201
Goalw [dvd_def] "!! a::'a::ring. a dvd 0";
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10789
diff changeset
   202
by (res_inst_tac [("x", "0")] exI 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   203
by (Simp_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   204
qed "dvd_zero_right";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   205
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10789
diff changeset
   206
Goalw [dvd_def] "!! a::'a::ring. 0 dvd a ==> a = 0";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   207
by Auto_tac;
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   208
qed "dvd_zero_left";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   209
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   210
Goalw [dvd_def] "!! a::'a::ring. a dvd a";
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   211
by (res_inst_tac [("x", "1")] exI 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   212
by (Simp_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   213
qed "dvd_refl_ring";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   214
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   215
Goalw [dvd_def] "!! a::'a::ring. [| a dvd b; b dvd c |] ==> a dvd c";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   216
by (Step_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   217
by (res_inst_tac [("x", "k * ka")] exI 1);
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   218
by (Asm_simp_tac 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   219
qed "dvd_trans_ring";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   220
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   221
Addsimps [dvd_zero_right, dvd_refl_ring];
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   222
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   223
Goalw [dvd_def]
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   224
  "!!a::'a::ring. [| a dvd 1; b dvd 1 |] ==> a * b dvd 1";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   225
by (Clarify_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   226
by (res_inst_tac [("x", "k * ka")] exI 1);
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   227
by (Asm_full_simp_tac 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   228
qed "unit_mult";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   229
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   230
Goal "!!a::'a::ring. a dvd 1 ==> a^n dvd 1";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   231
by (induct_tac "n" 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   232
by (Simp_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   233
by (asm_simp_tac (simpset() addsimps [unit_mult]) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   234
qed "unit_power";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   235
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   236
Goalw [dvd_def]
10789
260fa2c67e3e Changed priority of dvd from 70 to 50 as befits a relation.
nipkow
parents: 10448
diff changeset
   237
  "!! a::'a::ring. [| a dvd b; a dvd c |] ==> a dvd b + c";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   238
by (Clarify_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   239
by (res_inst_tac [("x", "k + ka")] exI 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   240
by (simp_tac (simpset() addsimps [r_distr]) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   241
qed "dvd_add_right";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   242
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   243
Goalw [dvd_def]
10789
260fa2c67e3e Changed priority of dvd from 70 to 50 as befits a relation.
nipkow
parents: 10448
diff changeset
   244
  "!! a::'a::ring. a dvd b ==> a dvd -b";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   245
by (Clarify_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   246
by (res_inst_tac [("x", "-k")] exI 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   247
by (simp_tac (simpset() addsimps [r_minus]) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   248
qed "dvd_uminus_right";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   249
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   250
Goalw [dvd_def]
10789
260fa2c67e3e Changed priority of dvd from 70 to 50 as befits a relation.
nipkow
parents: 10448
diff changeset
   251
  "!! a::'a::ring. a dvd b ==> a dvd c*b";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   252
by (Clarify_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   253
by (res_inst_tac [("x", "c * k")] exI 1);
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   254
by (Simp_tac 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   255
qed "dvd_l_mult_right";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   256
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   257
Goalw [dvd_def]
10789
260fa2c67e3e Changed priority of dvd from 70 to 50 as befits a relation.
nipkow
parents: 10448
diff changeset
   258
  "!! a::'a::ring. a dvd b ==> a dvd b*c";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   259
by (Clarify_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   260
by (res_inst_tac [("x", "k * c")] exI 1);
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   261
by (Simp_tac 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   262
qed "dvd_r_mult_right";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   263
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   264
Addsimps [dvd_add_right, dvd_uminus_right, dvd_l_mult_right, dvd_r_mult_right];
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   265
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   266
(* Inverse of multiplication *)
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   267
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   268
section "inverse";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   269
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   270
Goal "!! a::'a::ring. [| a * x = 1; a * y = 1 |] ==> x = y";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   271
by (res_inst_tac [("a", "(a*y)*x"), ("b", "y*(a*x)")] box_equals 1);
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   272
by (Simp_tac 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   273
by Auto_tac;
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   274
qed "inverse_unique";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   275
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   276
Goal "!! a::'a::ring. a dvd 1 ==> a * inverse a = 1";
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   277
by (asm_full_simp_tac (simpset () addsimps [inverse_def, dvd_def]) 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   278
by (Clarify_tac 1);
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   279
by (rtac theI 1);
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   280
by (atac 1);
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   281
by (rtac inverse_unique 1);
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   282
by (atac 1);
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   283
by (atac 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   284
qed "r_inverse_ring";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   285
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   286
Goal "!! a::'a::ring. a dvd 1 ==> inverse a * a= 1";
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   287
by (asm_simp_tac (simpset() addsimps [r_inverse_ring]) 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   288
qed "l_inverse_ring";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   289
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   290
(* Integral domain *)
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   291
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   292
(*
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   293
section "Integral domains";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   294
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   295
Goal "0 ~= (1::'a::domain)";
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10789
diff changeset
   296
by (rtac not_sym 1);
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10789
diff changeset
   297
by (rtac one_not_zero 1);
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10789
diff changeset
   298
qed "zero_not_one";
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10789
diff changeset
   299
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   300
Goal "!! a::'a::domain. a dvd 1 ==> a ~= 0";
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10789
diff changeset
   301
by (auto_tac (claset() addDs [dvd_zero_left],
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10789
diff changeset
   302
  simpset() addsimps [one_not_zero] ));
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10789
diff changeset
   303
qed "unit_imp_nonzero";
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10789
diff changeset
   304
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10789
diff changeset
   305
Goal "[| a * b = 0; a ~= 0 |] ==> (b::'a::domain) = 0";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   306
by (dtac integral 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   307
by (Fast_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   308
qed "r_integral";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   309
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10789
diff changeset
   310
Goal "[| a * b = 0; b ~= 0 |] ==> (a::'a::domain) = 0";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   311
by (dtac integral 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   312
by (Fast_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   313
qed "l_integral";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   314
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10789
diff changeset
   315
Goal "!! a::'a::domain. [| a ~= 0; b ~= 0 |] ==> a * b ~= 0";
10230
5eb935d6cc69 tidying and renaming of contrapos rules
paulson
parents: 10198
diff changeset
   316
by (blast_tac (claset() addIs [l_integral]) 1); 
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   317
qed "not_integral";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   318
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10789
diff changeset
   319
Addsimps [not_integral, one_not_zero, zero_not_one];
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   320
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   321
Goal "!! a::'a::domain. [| a * x = x; x ~= 0 |] ==> a = 1";
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   322
by (res_inst_tac [("a", "- 1")] a_lcancel 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   323
by (Simp_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   324
by (rtac l_integral 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   325
by (assume_tac 2);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   326
by (asm_simp_tac (simpset() addsimps [l_distr, l_minus]) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   327
qed "l_one_integral";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   328
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   329
Goal "!! a::'a::domain. [| x * a = x; x ~= 0 |] ==> a = 1";
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   330
by (res_inst_tac [("a", "- 1")] a_rcancel 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   331
by (Simp_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   332
by (rtac r_integral 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   333
by (assume_tac 2);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   334
by (asm_simp_tac (simpset() addsimps [r_distr, r_minus]) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   335
qed "r_one_integral";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   336
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   337
(* cancellation laws for multiplication *)
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   338
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10789
diff changeset
   339
Goal "!! a::'a::domain. [| a ~= 0; a * b = a * c |] ==> b = c";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   340
by (rtac diff_zero_imp_eq 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   341
by (dtac eq_imp_diff_zero 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   342
by (full_simp_tac (simpset() addsimps [r_minus RS sym, r_distr RS sym]) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   343
by (fast_tac (claset() addIs [l_integral]) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   344
qed "m_lcancel";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   345
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10789
diff changeset
   346
Goal "!! a::'a::domain. [| a ~= 0; b * a = c * a |] ==> b = c";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   347
by (rtac m_lcancel 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   348
by (assume_tac 1);
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   349
by (Asm_full_simp_tac 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   350
qed "m_rcancel";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   351
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10789
diff changeset
   352
Goal "!! a::'a::domain. a ~= 0 ==> (a * b = a * c) = (b = c)";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   353
by (auto_tac (claset() addDs [m_lcancel], simpset()));
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   354
qed "m_lcancel_eq";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   355
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10789
diff changeset
   356
Goal "!! a::'a::domain. a ~= 0 ==> (b * a = c * a) = (b = c)";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   357
by (asm_simp_tac (simpset() addsimps [m_lcancel_eq, m_comm]) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   358
qed "m_rcancel_eq";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   359
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   360
Addsimps [m_lcancel_eq, m_rcancel_eq];
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   361
*)
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   362
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   363
(* Fields *)
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   364
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   365
section "Fields";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   366
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   367
Goal "!! a::'a::field. (a dvd 1) = (a ~= 0)";
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   368
by (auto_tac (claset() addDs [thm "field_ax", dvd_zero_left],
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   369
  simpset() addsimps [thm "field_one_not_zero"]));
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   370
qed "field_unit";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   371
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   372
(* required for instantiation of field < domain in file Field.thy *)
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   373
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   374
Addsimps [field_unit];
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   375
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   376
val field_one_not_zero = thm "field_one_not_zero";
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   377
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   378
Goal "!! a::'a::field. a ~= 0 ==> a * inverse a = 1";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   379
by (asm_full_simp_tac (simpset() addsimps [r_inverse_ring]) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   380
qed "r_inverse";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   381
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   382
Goal "!! a::'a::field. a ~= 0 ==> inverse a * a= 1";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   383
by (asm_full_simp_tac (simpset() addsimps [l_inverse_ring]) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   384
qed "l_inverse";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   385
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   386
Addsimps [l_inverse, r_inverse];
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   387
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10789
diff changeset
   388
(* fields are integral domains *)
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   389
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10789
diff changeset
   390
Goal "!! a::'a::field. a * b = 0 ==> a = 0 | b = 0";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   391
by (Step_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   392
by (res_inst_tac [("a", "(a*b)*inverse b")] box_equals 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   393
by (rtac refl 3);
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   394
by (Simp_tac 2);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   395
by Auto_tac;
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   396
qed "field_integral";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   397
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10789
diff changeset
   398
(* fields are factorial domains *)
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10789
diff changeset
   399
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   400
Goalw [thm "prime_def", thm "irred_def"]
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   401
  "!! a::'a::field. irred a ==> prime a";
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13550
diff changeset
   402
by (blast_tac (claset() addIs [thm "field_ax"]) 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   403
qed "field_fact_prime";