src/HOL/Algebra/CRing.thy
author ballarin
Fri, 14 Mar 2003 18:00:16 +0100
changeset 13864 f44f121dd275
parent 13854 91c9ab25fece
child 13889 6676ac2527fa
permissions -rw-r--r--
Bugs fixed and operators finprod and finsum.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13835
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
     1
(*
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
     2
  Title:     The algebraic hierarchy of rings
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
     3
  Id:        $Id$
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
     4
  Author:    Clemens Ballarin, started 9 December 1996
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
     5
  Copyright: Clemens Ballarin
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
     6
*)
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
     7
13864
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
     8
theory CRing = Summation
13854
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
     9
files ("ringsimp.ML"):
13835
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    10
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    11
section {* The Algebraic Hierarchy of Rings *}
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    12
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    13
subsection {* Basic Definitions *}
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    14
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    15
record 'a ring = "'a group" +
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    16
  zero :: 'a ("\<zero>\<index>")
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    17
  add :: "['a, 'a] => 'a" (infixl "\<oplus>\<index>" 65)
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    18
  a_inv :: "'a => 'a" ("\<ominus>\<index> _" [81] 80)
13854
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
    19
  minus :: "['a, 'a] => 'a" (infixl "\<ominus>\<index>" 65)
13835
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    20
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    21
locale cring = abelian_monoid R +
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    22
  assumes a_abelian_group: "abelian_group (| carrier = carrier R,
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    23
      mult = add R, one = zero R, m_inv = a_inv R |)"
13854
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
    24
    and minus_def: "[| x \<in> carrier R; y \<in> carrier R |] ==> x \<ominus> y = x \<oplus> \<ominus> y"
13835
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    25
    and m_inv_def: "[| EX y. y \<in> carrier R & x \<otimes> y = \<one>; x \<in> carrier R |]
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    26
      ==> inv x = (THE y. y \<in> carrier R & x \<otimes> y = \<one>)"
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    27
    and l_distr: "[| x \<in> carrier R; y \<in> carrier R; z \<in> carrier R |]
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    28
      ==> (x \<oplus> y) \<otimes> z = x \<otimes> z \<oplus> y \<otimes> z"
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    29
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    30
(*
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    31
  -- {* Definition of derived operations *}
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    32
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    33
  minus_def:    "a - b = a + (-b)"
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    34
  inverse_def:  "inverse a = (if a dvd 1 then THE x. a*x = 1 else 0)"
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    35
  divide_def:   "a / b = a * inverse b"
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    36
  power_def:    "a ^ n = nat_rec 1 (%u b. b * a) n"
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    37
*)
13864
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
    38
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
    39
locale "domain" = cring +
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
    40
  assumes one_not_zero [simp]: "\<one> ~= \<zero>"
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
    41
    and integral: "[| a \<otimes> b = \<zero>; a \<in> carrier R; b \<in> carrier R |] ==>
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
    42
                  a = \<zero> | b = \<zero>"
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
    43
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
    44
subsection {* Basic Facts of Rings *}
13835
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    45
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    46
lemma (in cring) a_magma [simp, intro]:
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    47
  "magma (| carrier = carrier R,
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    48
     mult = add R, one = zero R, m_inv = a_inv R |)"
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    49
  using a_abelian_group by (simp only: abelian_group_def)
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    50
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    51
lemma (in cring) a_l_one [simp, intro]:
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    52
  "l_one (| carrier = carrier R,
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    53
     mult = add R, one = zero R, m_inv = a_inv R |)"
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    54
  using a_abelian_group by (simp only: abelian_group_def)
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    55
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    56
lemma (in cring) a_abelian_group_parts [simp, intro]:
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    57
  "semigroup_axioms (| carrier = carrier R,
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    58
     mult = add R, one = zero R, m_inv = a_inv R |)"
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    59
  "group_axioms (| carrier = carrier R,
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    60
     mult = add R, one = zero R, m_inv = a_inv R |)"
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    61
  "abelian_semigroup_axioms (| carrier = carrier R,
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    62
     mult = add R, one = zero R, m_inv = a_inv R |)"
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    63
  using a_abelian_group by (simp_all only: abelian_group_def)
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    64
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    65
lemma (in cring) a_semigroup:
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    66
  "semigroup (| carrier = carrier R,
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    67
     mult = add R, one = zero R, m_inv = a_inv R |)"
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    68
  by (simp add: semigroup_def)
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    69
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    70
lemma (in cring) a_group:
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    71
  "group (| carrier = carrier R,
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    72
     mult = add R, one = zero R, m_inv = a_inv R |)"
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    73
  by (simp add: group_def)
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    74
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    75
lemma (in cring) a_abelian_semigroup:
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    76
  "abelian_semigroup (| carrier = carrier R,
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    77
     mult = add R, one = zero R, m_inv = a_inv R |)"
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    78
  by (simp add: abelian_semigroup_def)
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    79
13864
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
    80
lemmas group_record_simps = semigroup.simps monoid.simps group.simps
13835
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    81
13854
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
    82
lemmas (in cring) a_closed [intro, simp] =
13864
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
    83
  magma.m_closed [OF a_magma, simplified group_record_simps]
13854
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
    84
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
    85
lemmas (in cring) zero_closed [intro, simp] = 
13864
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
    86
  l_one.one_closed [OF a_l_one, simplified group_record_simps]
13854
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
    87
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
    88
lemmas (in cring) a_inv_closed [intro, simp] =
13864
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
    89
  group.inv_closed [OF a_group, simplified group_record_simps]
13854
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
    90
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
    91
lemma (in cring) minus_closed [intro, simp]:
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
    92
  "[| x \<in> carrier R; y \<in> carrier R |] ==> x \<ominus> y \<in> carrier R"
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
    93
  by (simp add: minus_def)
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
    94
13864
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
    95
lemmas (in cring) a_l_cancel [simp] =
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
    96
  group.l_cancel [OF a_group, simplified group_record_simps]
13835
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
    97
13864
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
    98
lemmas (in cring) a_r_cancel [simp] =
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
    99
  group.r_cancel [OF a_group, simplified group_record_simps]
13854
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   100
13864
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   101
lemmas (in cring) a_assoc =
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   102
  semigroup.m_assoc [OF a_semigroup, simplified group_record_simps]
13854
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   103
13864
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   104
lemmas (in cring) l_zero [simp] =
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   105
  l_one.l_one [OF a_l_one, simplified group_record_simps]
13854
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   106
13864
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   107
lemmas (in cring) l_neg =
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   108
  group.l_inv [OF a_group, simplified group_record_simps]
13854
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   109
13864
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   110
lemmas (in cring) a_comm =
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   111
  abelian_semigroup.m_comm [OF a_abelian_semigroup,
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   112
  simplified group_record_simps]
13854
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   113
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   114
lemmas (in cring) a_lcomm =
13864
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   115
  abelian_semigroup.m_lcomm [OF a_abelian_semigroup, simplified group_record_simps]
13854
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   116
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   117
lemma (in cring) r_zero [simp]:
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   118
  "x \<in> carrier R ==> x \<oplus> \<zero> = x"
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   119
  using group.r_one [OF a_group]
13835
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
   120
  by simp
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
   121
13854
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   122
lemma (in cring) r_neg:
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   123
  "x \<in> carrier R ==> x \<oplus> (\<ominus> x) = \<zero>"
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   124
  using group.r_inv [OF a_group]
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   125
  by simp
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   126
13864
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   127
lemmas (in cring) minus_zero [simp] =
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   128
  group.inv_one [OF a_group, simplified group_record_simps]
13854
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   129
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   130
lemma (in cring) minus_minus [simp]:
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   131
  "x \<in> carrier R ==> \<ominus> (\<ominus> x) = x"
13864
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   132
  using group.inv_inv [OF a_group, simplified group_record_simps]
13835
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
   133
  by simp
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
   134
13854
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   135
lemma (in cring) minus_add:
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   136
  "[| x \<in> carrier R; y \<in> carrier R |] ==> \<ominus> (x \<oplus> y) = \<ominus> x \<oplus> \<ominus> y"
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   137
  using abelian_group.inv_mult [OF a_abelian_group]
13835
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
   138
  by simp
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
   139
13854
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   140
lemmas (in cring) a_ac = a_assoc a_comm a_lcomm
13835
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
   141
13854
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   142
subsection {* Normaliser for Commutative Rings *}
13835
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
   143
13854
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   144
lemma (in cring) r_neg2:
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   145
  "[| x \<in> carrier R; y \<in> carrier R |] ==> x \<oplus> (\<ominus> x \<oplus> y) = y"
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   146
proof -
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   147
  assume G: "x \<in> carrier R" "y \<in> carrier R"
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   148
  then have "(x \<oplus> \<ominus> x) \<oplus> y = y" by (simp only: r_neg l_zero)
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   149
  with G show ?thesis by (simp add: a_ac)
13835
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
   150
qed
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
   151
13854
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   152
lemma (in cring) r_neg1:
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   153
  "[| x \<in> carrier R; y \<in> carrier R |] ==> \<ominus> x \<oplus> (x \<oplus> y) = y"
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   154
proof -
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   155
  assume G: "x \<in> carrier R" "y \<in> carrier R"
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   156
  then have "(\<ominus> x \<oplus> x) \<oplus> y = y" by (simp only: l_neg l_zero)
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   157
  with G show ?thesis by (simp add: a_ac)
13835
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
   158
qed
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
   159
13854
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   160
lemma (in cring) r_distr:
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   161
  "[| x \<in> carrier R; y \<in> carrier R; z \<in> carrier R |]
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   162
  ==> z \<otimes> (x \<oplus> y) = z \<otimes> x \<oplus> z \<otimes> y"
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   163
proof -
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   164
  assume G: "x \<in> carrier R" "y \<in> carrier R" "z \<in> carrier R"
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   165
  then have "z \<otimes> (x \<oplus> y) = (x \<oplus> y) \<otimes> z" by (simp add: m_comm)
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   166
  also from G have "... = x \<otimes> z \<oplus> y \<otimes> z" by (simp add: l_distr)
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   167
  also from G have "... = z \<otimes> x \<oplus> z \<otimes> y" by (simp add: m_comm)
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   168
  finally show ?thesis .
13835
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
   169
qed
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
   170
13854
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   171
text {* 
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   172
  The following proofs are from Jacobson, Basic Algebra I, pp.~88--89
13835
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
   173
*}
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
   174
13854
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   175
lemma (in cring) l_null [simp]:
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   176
  "x \<in> carrier R ==> \<zero> \<otimes> x = \<zero>"
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   177
proof -
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   178
  assume R: "x \<in> carrier R"
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   179
  then have "\<zero> \<otimes> x \<oplus> \<zero> \<otimes> x = (\<zero> \<oplus> \<zero>) \<otimes> x"
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   180
    by (simp add: l_distr del: l_zero r_zero)
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   181
  also from R have "... = \<zero> \<otimes> x \<oplus> \<zero>" by simp
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   182
  finally have "\<zero> \<otimes> x \<oplus> \<zero> \<otimes> x = \<zero> \<otimes> x \<oplus> \<zero>" .
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   183
  with R show ?thesis by (simp del: r_zero)
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   184
qed
13835
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
   185
13854
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   186
lemma (in cring) r_null [simp]:
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   187
  "x \<in> carrier R ==> x \<otimes> \<zero> = \<zero>"
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   188
proof -
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   189
  assume R: "x \<in> carrier R"
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   190
  then have "x \<otimes> \<zero> = \<zero> \<otimes> x" by (simp add: ac)
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   191
  also from R have "... = \<zero>" by simp
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   192
  finally show ?thesis .
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   193
qed
13835
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
   194
13854
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   195
lemma (in cring) l_minus:
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   196
  "[| x \<in> carrier R; y \<in> carrier R |] ==> \<ominus> x \<otimes> y = \<ominus> (x \<otimes> y)"
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   197
proof -
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   198
  assume R: "x \<in> carrier R" "y \<in> carrier R"
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   199
  then have "(\<ominus> x) \<otimes> y \<oplus> x \<otimes> y = (\<ominus> x \<oplus> x) \<otimes> y" by (simp add: l_distr)
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   200
  also from R have "... = \<zero>" by (simp add: l_neg l_null)
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   201
  finally have "(\<ominus> x) \<otimes> y \<oplus> x \<otimes> y = \<zero>" .
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   202
  with R have "(\<ominus> x) \<otimes> y \<oplus> x \<otimes> y \<oplus> \<ominus> (x \<otimes> y) = \<zero> \<oplus> \<ominus> (x \<otimes> y)" by simp
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   203
  with R show ?thesis by (simp add: a_assoc r_neg )
13835
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
   204
qed
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
   205
13854
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   206
lemma (in cring) r_minus:
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   207
  "[| x \<in> carrier R; y \<in> carrier R |] ==> x \<otimes> \<ominus> y = \<ominus> (x \<otimes> y)"
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   208
proof -
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   209
  assume R: "x \<in> carrier R" "y \<in> carrier R"
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   210
  then have "x \<otimes> \<ominus> y = \<ominus> y \<otimes> x" by (simp add: ac)
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   211
  also from R have "... = \<ominus> (y \<otimes> x)" by (simp add: l_minus)
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   212
  also from R have "... = \<ominus> (x \<otimes> y)" by (simp add: ac)
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   213
  finally show ?thesis .
13835
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
   214
qed
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
   215
13854
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   216
lemmas (in cring) cring_simprules =
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   217
  a_closed zero_closed a_inv_closed minus_closed m_closed one_closed
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   218
  a_assoc l_zero l_neg a_comm m_assoc l_one l_distr m_comm minus_def
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   219
  r_zero r_neg r_neg2 r_neg1 minus_add minus_minus minus_zero
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   220
  a_lcomm m_lcomm r_distr l_null r_null l_minus r_minus
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   221
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   222
use "ringsimp.ML"
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   223
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   224
method_setup algebra =
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   225
  {* Method.ctxt_args cring_normalise *}
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   226
  {* computes distributive normal form in commutative rings (locales version) *}
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   227
13864
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   228
text {* Two examples for use of method algebra *}
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   229
13854
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   230
lemma
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   231
  includes cring R + cring S
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   232
  shows "[| a \<in> carrier R; b \<in> carrier R; c \<in> carrier S; d \<in> carrier S |] ==> 
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   233
  a \<oplus> \<ominus> (a \<oplus> \<ominus> b) = b & c \<otimes>\<^sub>2 d = d \<otimes>\<^sub>2 c"
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   234
  by algebra
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   235
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   236
lemma
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   237
  includes cring
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   238
  shows "[| a \<in> carrier R; b \<in> carrier R |] ==> a \<ominus> (a \<ominus> b) = b"
91c9ab25fece First distributed version of Group and Ring theory.
ballarin
parents: 13835
diff changeset
   239
  by algebra
13835
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
   240
13864
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   241
subsection {* Sums over Finite Sets *}
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   242
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   243
text {*
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   244
  This definition makes it easy to lift lemmas from @{term finprod}.
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   245
*}
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   246
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   247
constdefs
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   248
  finsum :: "[('b, 'm) ring_scheme, 'a => 'b, 'a set] => 'b"
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   249
  "finsum R f A == finprod (| carrier = carrier R,
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   250
     mult = add R, one = zero R, m_inv = a_inv R |) f A"
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   251
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   252
lemma (in cring) a_abelian_monoid:
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   253
  "abelian_monoid (| carrier = carrier R,
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   254
     mult = add R, one = zero R, m_inv = a_inv R |)"
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   255
  by (simp add: abelian_monoid_def)
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   256
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   257
(*
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   258
  lemmas (in cring) finsum_empty [simp] =
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   259
    abelian_monoid.finprod_empty [OF a_abelian_monoid, simplified]
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   260
  is dangeous, because attributes (like simplified) are applied upon opening
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   261
  the locale, simplified refers to the simpset at that time!!!
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   262
*)
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   263
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   264
lemmas (in cring) finsum_empty [simp] =
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   265
  abelian_monoid.finprod_empty [OF a_abelian_monoid, folded finsum_def,
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   266
    simplified group_record_simps]
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   267
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   268
lemmas (in cring) finsum_insert [simp] =
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   269
  abelian_monoid.finprod_insert [OF a_abelian_monoid, folded finsum_def,
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   270
    simplified group_record_simps]
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   271
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   272
lemmas (in cring) finsum_zero =
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   273
  abelian_monoid.finprod_one [OF a_abelian_monoid, folded finsum_def,
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   274
    simplified group_record_simps]
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   275
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   276
lemmas (in cring) finsum_closed [simp] =
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   277
  abelian_monoid.finprod_closed [OF a_abelian_monoid, folded finsum_def,
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   278
    simplified group_record_simps]
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   279
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   280
lemmas (in cring) finsum_Un_Int =
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   281
  abelian_monoid.finprod_Un_Int [OF a_abelian_monoid, folded finsum_def,
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   282
    simplified group_record_simps]
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   283
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   284
lemmas (in cring) finsum_Un_disjoint =
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   285
  abelian_monoid.finprod_Un_disjoint [OF a_abelian_monoid, folded finsum_def,
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   286
    simplified group_record_simps]
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   287
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   288
lemmas (in cring) finsum_addf =
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   289
  abelian_monoid.finprod_multf [OF a_abelian_monoid, folded finsum_def,
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   290
    simplified group_record_simps]
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   291
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   292
lemmas (in cring) finsum_cong =
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   293
  abelian_monoid.finprod_cong [OF a_abelian_monoid, folded finsum_def,
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   294
    simplified group_record_simps]
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   295
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   296
lemmas (in cring) finsum_0 [simp] =
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   297
  abelian_monoid.finprod_0 [OF a_abelian_monoid, folded finsum_def,
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   298
    simplified group_record_simps]
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   299
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   300
lemmas (in cring) finsum_Suc [simp] =
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   301
  abelian_monoid.finprod_Suc [OF a_abelian_monoid, folded finsum_def,
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   302
    simplified group_record_simps]
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   303
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   304
lemmas (in cring) finsum_Suc2 =
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   305
  abelian_monoid.finprod_Suc2 [OF a_abelian_monoid, folded finsum_def,
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   306
    simplified group_record_simps]
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   307
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   308
lemmas (in cring) finsum_add [simp] =
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   309
  abelian_monoid.finprod_mult [OF a_abelian_monoid, folded finsum_def,
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   310
    simplified group_record_simps]
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   311
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   312
lemmas (in cring) finsum_cong' [cong] =
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   313
  abelian_monoid.finprod_cong' [OF a_abelian_monoid, folded finsum_def,
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   314
    simplified group_record_simps]
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   315
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   316
(*
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   317
lemma (in cring) finsum_empty [simp]:
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   318
  "finsum R f {} = \<zero>"
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   319
  by (simp add: finsum_def
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   320
    abelian_monoid.finprod_empty [OF a_abelian_monoid])
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   321
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   322
lemma (in cring) finsum_insert [simp]:
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   323
  "[| finite F; a \<notin> F; f : F -> carrier R; f a \<in> carrier R |] ==>
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   324
   finsum R f (insert a F) = f a \<oplus> finsum R f F"
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   325
  by (simp add: finsum_def
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   326
    abelian_monoid.finprod_insert [OF a_abelian_monoid])
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   327
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   328
lemma (in cring) finsum_zero:
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   329
  "finite A ==> finsum R (%i. \<zero>) A = \<zero>"
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   330
  by (simp add: finsum_def
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   331
    abelian_monoid.finprod_one [OF a_abelian_monoid, simplified])
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   332
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   333
lemma (in cring) finsum_closed [simp]:
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   334
  "[| finite A; f : A -> carrier R |] ==> finsum R f A \<in> carrier R"
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   335
  by (simp only: finsum_def
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   336
    abelian_monoid.finprod_closed [OF a_abelian_monoid, simplified])
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   337
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   338
lemma (in cring) finsum_Un_Int:
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   339
  "[| finite A; finite B; g \<in> A -> carrier R; g \<in> B -> carrier R |] ==>
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   340
     finsum R g (A Un B) \<oplus> finsum R g (A Int B) =
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   341
     finsum R g A \<oplus> finsum R g B"
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   342
  by (simp only: finsum_def
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   343
    abelian_monoid.finprod_Un_Int [OF a_abelian_monoid, simplified])
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   344
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   345
lemma (in cring) finsum_Un_disjoint:
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   346
  "[| finite A; finite B; A Int B = {};
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   347
      g \<in> A -> carrier R; g \<in> B -> carrier R |]
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   348
   ==> finsum R g (A Un B) = finsum R g A \<oplus> finsum R g B"
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   349
  by (simp only: finsum_def
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   350
    abelian_monoid.finprod_Un_disjoint [OF a_abelian_monoid, simplified])
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   351
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   352
lemma (in cring) finsum_addf:
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   353
  "[| finite A; f \<in> A -> carrier R; g \<in> A -> carrier R |] ==>
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   354
   finsum R (%x. f x \<oplus> g x) A = (finsum R f A \<oplus> finsum R g A)"
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   355
  by (simp only: finsum_def
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   356
    abelian_monoid.finprod_multf [OF a_abelian_monoid, simplified])
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   357
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   358
lemma (in cring) finsum_cong:
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   359
  "[| A = B; g : B -> carrier R;
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   360
      !!i. i : B ==> f i = g i |] ==> finsum R f A = finsum R g B"
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   361
  apply (simp only: finsum_def)
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   362
  apply (rule abelian_monoid.finprod_cong [OF a_abelian_monoid, simplified])
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   363
  apply simp_all
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   364
  done
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   365
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   366
lemma (in cring) finsum_0 [simp]:
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   367
  "f \<in> {0::nat} -> carrier R ==> finsum R f {..0} = f 0"
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   368
  by (simp add: finsum_def
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   369
    abelian_monoid.finprod_0 [OF a_abelian_monoid, simplified])
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   370
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   371
lemma (in cring) finsum_Suc [simp]:
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   372
  "f \<in> {..Suc n} -> carrier R ==>
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   373
   finsum R f {..Suc n} = (f (Suc n) \<oplus> finsum R f {..n})"
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   374
  by (simp add: finsum_def
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   375
    abelian_monoid.finprod_Suc [OF a_abelian_monoid, simplified])
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   376
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   377
lemma (in cring) finsum_Suc2:
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   378
  "f \<in> {..Suc n} -> carrier R ==>
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   379
   finsum R f {..Suc n} = (finsum R (%i. f (Suc i)) {..n} \<oplus> f 0)"
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   380
  by (simp only: finsum_def
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   381
    abelian_monoid.finprod_Suc2 [OF a_abelian_monoid, simplified])
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   382
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   383
lemma (in cring) finsum_add [simp]:
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   384
  "[| f : {..n} -> carrier R; g : {..n} -> carrier R |] ==>
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   385
     finsum R (%i. f i \<oplus> g i) {..n::nat} =
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   386
     finsum R f {..n} \<oplus> finsum R g {..n}"
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   387
  by (simp only: finsum_def
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   388
    abelian_monoid.finprod_mult [OF a_abelian_monoid, simplified])
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   389
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   390
lemma (in cring) finsum_cong' [cong]:
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   391
  "[| A = B; !!i. i : B ==> f i = g i;
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   392
      g \<in> B -> carrier R = True |] ==> finsum R f A = finsum R g B"
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   393
  apply (simp only: finsum_def)
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   394
  apply (rule abelian_monoid.finprod_cong' [OF a_abelian_monoid, simplified])
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   395
  apply simp_all
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   396
  done
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   397
*)
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   398
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   399
text {*Usually, if this rule causes a failed congruence proof error,
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   400
   the reason is that the premise @{text "g \<in> B -> carrier G"} cannot be shown.
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   401
   Adding @{thm [source] Pi_def} to the simpset is often useful. *}
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   402
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   403
lemma (in cring) finsum_ldistr:
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   404
  "[| finite A; a \<in> carrier R; f \<in> A -> carrier R |] ==>
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   405
   finsum R f A \<otimes> a = finsum R (%i. f i \<otimes> a) A"
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   406
proof (induct set: Finites)
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   407
  case empty then show ?case by simp
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   408
next
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   409
  case (insert F x) then show ?case by (simp add: Pi_def l_distr)
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   410
qed
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   411
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   412
lemma (in cring) finsum_rdistr:
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   413
  "[| finite A; a \<in> carrier R; f \<in> A -> carrier R |] ==>
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   414
   a \<otimes> finsum R f A = finsum R (%i. a \<otimes> f i) A"
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   415
proof (induct set: Finites)
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   416
  case empty then show ?case by simp
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   417
next
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   418
  case (insert F x) then show ?case by (simp add: Pi_def r_distr)
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   419
qed
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   420
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   421
subsection {* Facts of Integral Domains *}
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   422
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   423
lemma (in "domain") zero_not_one [simp]:
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   424
  "\<zero> ~= \<one>"
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   425
  by (rule not_sym) simp
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   426
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   427
lemma (in "domain") integral_iff: (* not by default a simp rule! *)
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   428
  "[| a \<in> carrier R; b \<in> carrier R |] ==> (a \<otimes> b = \<zero>) = (a = \<zero> | b = \<zero>)"
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   429
proof
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   430
  assume "a \<in> carrier R" "b \<in> carrier R" "a \<otimes> b = \<zero>"
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   431
  then show "a = \<zero> | b = \<zero>" by (simp add: integral)
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   432
next
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   433
  assume "a \<in> carrier R" "b \<in> carrier R" "a = \<zero> | b = \<zero>"
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   434
  then show "a \<otimes> b = \<zero>" by auto
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   435
qed
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   436
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   437
lemma (in "domain") m_lcancel:
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   438
  assumes prem: "a ~= \<zero>"
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   439
    and R: "a \<in> carrier R" "b \<in> carrier R" "c \<in> carrier R"
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   440
  shows "(a \<otimes> b = a \<otimes> c) = (b = c)"
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   441
proof
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   442
  assume eq: "a \<otimes> b = a \<otimes> c"
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   443
  with R have "a \<otimes> (b \<ominus> c) = \<zero>" by algebra
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   444
  with R have "a = \<zero> | (b \<ominus> c) = \<zero>" by (simp add: integral_iff)
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   445
  with prem and R have "b \<ominus> c = \<zero>" by auto 
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   446
  with R have "b = b \<ominus> (b \<ominus> c)" by algebra 
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   447
  also from R have "b \<ominus> (b \<ominus> c) = c" by algebra
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   448
  finally show "b = c" .
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   449
next
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   450
  assume "b = c" then show "a \<otimes> b = a \<otimes> c" by simp
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   451
qed
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   452
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   453
lemma (in "domain") m_rcancel:
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   454
  assumes prem: "a ~= \<zero>"
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   455
    and R: "a \<in> carrier R" "b \<in> carrier R" "c \<in> carrier R"
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   456
  shows conc: "(b \<otimes> a = c \<otimes> a) = (b = c)"
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   457
proof -
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   458
  from prem and R have "(a \<otimes> b = a \<otimes> c) = (b = c)" by (rule m_lcancel)
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   459
  with R show ?thesis by algebra
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   460
qed
f44f121dd275 Bugs fixed and operators finprod and finsum.
ballarin
parents: 13854
diff changeset
   461
13835
12b2ffbe543a Change to meta simplifier: congruence rules may now have frees as head of term.
ballarin
parents:
diff changeset
   462
end