| author | wenzelm | 
| Tue, 16 Feb 2010 13:35:42 +0100 | |
| changeset 35144 | 8b8302da3a55 | 
| parent 35092 | cfe605c54e50 | 
| child 35267 | 8dfd816713c6 | 
| permissions | -rw-r--r-- | 
| 34974 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 1 | (* Title: HOL/Algebras.thy | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 2 | Author: Florian Haftmann, TU Muenchen | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 3 | *) | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 4 | |
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 5 | header {* Generic algebraic structures and operations *}
 | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 6 | |
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 7 | theory Algebras | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 8 | imports HOL | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 9 | begin | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 10 | |
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 11 | subsection {* Generic algebraic structures *}
 | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 12 | |
| 35026 
02850d0b95ac
added explaining comment; added ac_simps slot; drop unused abstract lattice; dropped mysterious syntax declaration
 haftmann parents: 
34974diff
changeset | 13 | text {*
 | 
| 
02850d0b95ac
added explaining comment; added ac_simps slot; drop unused abstract lattice; dropped mysterious syntax declaration
 haftmann parents: 
34974diff
changeset | 14 | These locales provide basic structures for interpretation into | 
| 
02850d0b95ac
added explaining comment; added ac_simps slot; drop unused abstract lattice; dropped mysterious syntax declaration
 haftmann parents: 
34974diff
changeset | 15 | bigger structures; extensions require careful thinking, otherwise | 
| 
02850d0b95ac
added explaining comment; added ac_simps slot; drop unused abstract lattice; dropped mysterious syntax declaration
 haftmann parents: 
34974diff
changeset | 16 | undesired effects may occur due to interpretation. | 
| 
02850d0b95ac
added explaining comment; added ac_simps slot; drop unused abstract lattice; dropped mysterious syntax declaration
 haftmann parents: 
34974diff
changeset | 17 | *} | 
| 
02850d0b95ac
added explaining comment; added ac_simps slot; drop unused abstract lattice; dropped mysterious syntax declaration
 haftmann parents: 
34974diff
changeset | 18 | |
| 
02850d0b95ac
added explaining comment; added ac_simps slot; drop unused abstract lattice; dropped mysterious syntax declaration
 haftmann parents: 
34974diff
changeset | 19 | ML {*
 | 
| 
02850d0b95ac
added explaining comment; added ac_simps slot; drop unused abstract lattice; dropped mysterious syntax declaration
 haftmann parents: 
34974diff
changeset | 20 | structure Ac_Simps = Named_Thms( | 
| 
02850d0b95ac
added explaining comment; added ac_simps slot; drop unused abstract lattice; dropped mysterious syntax declaration
 haftmann parents: 
34974diff
changeset | 21 | val name = "ac_simps" | 
| 
02850d0b95ac
added explaining comment; added ac_simps slot; drop unused abstract lattice; dropped mysterious syntax declaration
 haftmann parents: 
34974diff
changeset | 22 | val description = "associativity and commutativity simplification rules" | 
| 
02850d0b95ac
added explaining comment; added ac_simps slot; drop unused abstract lattice; dropped mysterious syntax declaration
 haftmann parents: 
34974diff
changeset | 23 | ) | 
| 
02850d0b95ac
added explaining comment; added ac_simps slot; drop unused abstract lattice; dropped mysterious syntax declaration
 haftmann parents: 
34974diff
changeset | 24 | *} | 
| 
02850d0b95ac
added explaining comment; added ac_simps slot; drop unused abstract lattice; dropped mysterious syntax declaration
 haftmann parents: 
34974diff
changeset | 25 | |
| 
02850d0b95ac
added explaining comment; added ac_simps slot; drop unused abstract lattice; dropped mysterious syntax declaration
 haftmann parents: 
34974diff
changeset | 26 | setup Ac_Simps.setup | 
| 
02850d0b95ac
added explaining comment; added ac_simps slot; drop unused abstract lattice; dropped mysterious syntax declaration
 haftmann parents: 
34974diff
changeset | 27 | |
| 34974 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 28 | locale semigroup = | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 29 | fixes f :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "*" 70) | 
| 35026 
02850d0b95ac
added explaining comment; added ac_simps slot; drop unused abstract lattice; dropped mysterious syntax declaration
 haftmann parents: 
34974diff
changeset | 30 | assumes assoc [ac_simps]: "a * b * c = a * (b * c)" | 
| 34974 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 31 | |
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 32 | locale abel_semigroup = semigroup + | 
| 35026 
02850d0b95ac
added explaining comment; added ac_simps slot; drop unused abstract lattice; dropped mysterious syntax declaration
 haftmann parents: 
34974diff
changeset | 33 | assumes commute [ac_simps]: "a * b = b * a" | 
| 34974 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 34 | begin | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 35 | |
| 35026 
02850d0b95ac
added explaining comment; added ac_simps slot; drop unused abstract lattice; dropped mysterious syntax declaration
 haftmann parents: 
34974diff
changeset | 36 | lemma left_commute [ac_simps]: | 
| 34974 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 37 | "b * (a * c) = a * (b * c)" | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 38 | proof - | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 39 | have "(b * a) * c = (a * b) * c" | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 40 | by (simp only: commute) | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 41 | then show ?thesis | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 42 | by (simp only: assoc) | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 43 | qed | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 44 | |
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 45 | end | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 46 | |
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 47 | locale semilattice = abel_semigroup + | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 48 | assumes idem [simp]: "a * a = a" | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 49 | begin | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 50 | |
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 51 | lemma left_idem [simp]: | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 52 | "a * (a * b) = a * b" | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 53 | by (simp add: assoc [symmetric]) | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 54 | |
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 55 | end | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 56 | |
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 57 | |
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35084diff
changeset | 58 | subsection {* Generic syntactic operations *}
 | 
| 34974 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 59 | |
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 60 | class zero = | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 61 |   fixes zero :: 'a  ("0")
 | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 62 | |
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 63 | class one = | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 64 |   fixes one  :: 'a  ("1")
 | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 65 | |
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35084diff
changeset | 66 | hide (open) const zero one | 
| 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35084diff
changeset | 67 | |
| 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35084diff
changeset | 68 | syntax | 
| 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35084diff
changeset | 69 |   "_index1"  :: index    ("\<^sub>1")
 | 
| 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35084diff
changeset | 70 | translations | 
| 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35084diff
changeset | 71 | (index) "\<^sub>1" => (index) "\<^bsub>\<struct>\<^esub>" | 
| 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35084diff
changeset | 72 | |
| 34974 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 73 | lemma Let_0 [simp]: "Let 0 f = f 0" | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 74 | unfolding Let_def .. | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 75 | |
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 76 | lemma Let_1 [simp]: "Let 1 f = f 1" | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 77 | unfolding Let_def .. | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 78 | |
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 79 | setup {*
 | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 80 | Reorient_Proc.add | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 81 |     (fn Const(@{const_name Algebras.zero}, _) => true
 | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 82 |       | Const(@{const_name Algebras.one}, _) => true
 | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 83 | | _ => false) | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 84 | *} | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 85 | |
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 86 | simproc_setup reorient_zero ("0 = x") = Reorient_Proc.proc
 | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 87 | simproc_setup reorient_one ("1 = x") = Reorient_Proc.proc
 | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 88 | |
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 89 | typed_print_translation {*
 | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 90 | let | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 91 | fun tr' c = (c, fn show_sorts => fn T => fn ts => | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 92 | if (not o null) ts orelse T = dummyT | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 93 | orelse not (! show_types) andalso can Term.dest_Type T | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 94 | then raise Match | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 95 | else Syntax.const Syntax.constrainC $ Syntax.const c $ Syntax.term_of_typ show_sorts T); | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 96 | in map tr' [@{const_syntax Algebras.one}, @{const_syntax Algebras.zero}] end;
 | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 97 | *} -- {* show types that are presumably too general *}
 | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 98 | |
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 99 | class plus = | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 100 | fixes plus :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "+" 65) | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 101 | |
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 102 | class minus = | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 103 | fixes minus :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "-" 65) | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 104 | |
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 105 | class uminus = | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 106 |   fixes uminus :: "'a \<Rightarrow> 'a"  ("- _" [81] 80)
 | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 107 | |
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 108 | class times = | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 109 | fixes times :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "*" 70) | 
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 110 | |
| 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 haftmann parents: diff
changeset | 111 | end |