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