src/HOL/GroupTheory/Module.thy
author paulson
Wed, 29 Jan 2003 16:34:51 +0100
changeset 13792 d1811693899c
parent 13589 b6d1a29dc978
permissions -rw-r--r--
converted more UNITY theories to new-style
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13589
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
     1
(*  Title:      HOL/GroupTheory/Module
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
     2
    ID:         $Id$
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
     4
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
     5
*)
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
     6
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
     7
header{*Modules*}
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
     8
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
     9
theory Module = Ring:
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    10
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    11
text{*The record includes all the group components (carrier, sum, gminus,
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    12
zero), adding the scalar product.*}
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    13
record ('a,'b) module  = "'a group" + 
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    14
  sprod :: "'b \<Rightarrow> 'a \<Rightarrow> 'a"   (infixl "\<star>\<index>" 70) 
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    15
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    16
text{*The locale assumes all the properties of a ring and an Abelian group,
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    17
adding laws relating them.*}
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    18
locale module = ring R + abelian_group M +
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    19
  assumes sprod_funcset: "sprod M \<in> carrier R \<rightarrow> carrier M \<rightarrow> carrier M"
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    20
      and sprod_assoc:   
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    21
            "[|r \<in> carrier R; s \<in> carrier R; a \<in> carrier M|] 
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    22
             ==> (r \<cdot>\<^sub>1 s) \<star>\<^sub>2 a = r \<star>\<^sub>2 (s \<star>\<^sub>2 a)"
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    23
      and sprod_distrib_left:
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    24
            "[|r \<in> carrier R; a \<in> carrier M; b \<in> carrier M|] 
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    25
             ==> r \<star>\<^sub>2 (a \<oplus>\<^sub>2 b) = (r \<star>\<^sub>2 a) \<oplus>\<^sub>2 (r \<star>\<^sub>2 b)"
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    26
      and sprod_distrib_right:
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    27
            "[|r \<in> carrier R; s \<in> carrier R; a \<in> carrier M|] 
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    28
             ==> (r \<oplus>\<^sub>1 s) \<star>\<^sub>2 a = (r \<star>\<^sub>2 a) \<oplus>\<^sub>2 (s \<star>\<^sub>2 a)"
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    29
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    30
lemma module_iff:
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    31
     "module R M = (ring R & abelian_group M & module_axioms R M)"
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    32
by (simp add: module_def ring_def abelian_group_def) 
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    33
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    34
text{*The sum and product in @{text R} are @{text "r \<oplus>\<^sub>1 s"} and @{text
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    35
"r \<cdot>\<^sub>1 s"}, respectively.  The sum in @{text M} is @{text "a \<oplus>\<^sub>2 b"}.*}
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    36
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    37
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    38
text{*We have to write the ring product as @{text "\<cdot>\<^sub>2"}. But if we
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    39
refer to the scalar product as @{text "\<cdot>\<^sub>1"} then syntactic ambiguities
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    40
arise.  This presumably happens because the subscript merely determines which
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    41
structure argument to insert, which happens after parsing.  Better to use a
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    42
different symbol such as @{text "\<star>"}*}
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    43
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    44
subsection {*Trivial Example: Every Ring is an R-Module Over Itself *}
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    45
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    46
constdefs
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    47
 triv_mod :: "('a,'b) ring_scheme => ('a,'a) module"
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    48
  "triv_mod R == (|carrier = carrier R,
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    49
                 sum = sum R,
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    50
                 gminus = gminus R,
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    51
                 zero = zero R, 
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    52
                 sprod = prod R|)"
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    53
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    54
theorem module_triv_mod: "ring R ==> module R (triv_mod R)"
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    55
apply (simp add: triv_mod_def module_iff module_axioms_def
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    56
                 ring_def ring_axioms_def abelian_group_def
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    57
                 distrib_l_def distrib_r_def semigroup_def group_axioms_def)
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    58
apply (simp add: abelian_group_axioms_def)
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    59
  --{*Combining the two simplifications causes looping!*}
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    60
done
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    61
b6d1a29dc978 New theory GroupTheory/Module.thy of modules
paulson
parents:
diff changeset
    62
end