src/HOL/ex/IntRing.thy
author paulson
Tue, 23 May 2000 18:08:52 +0200
changeset 8936 a1c426541757
parent 6921 78a2ce8fb8df
child 11701 3d51fbf81c17
permissions -rw-r--r--
Now that 0 is overloaded, constant "zero" and its type class "zero" are no longer needed. Also IntRingDefs is redundant
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     1
(*  Title:      HOL/Integ/IntRing.thy
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     2
    ID:         $Id$
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     3
    Author:     Tobias Nipkow and Markus Wenzel
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     4
    Copyright   1996 TU Muenchen
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     5
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     6
The integers form a commutative ring.
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     7
With an application of Lagrange's lemma.
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     8
*)
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     9
8936
a1c426541757 Now that 0 is overloaded, constant "zero" and its type class "zero" are
paulson
parents: 6921
diff changeset
    10
IntRing = Ring + Lagrange +
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    11
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    12
instance int :: add_semigroup (zadd_assoc)
8936
a1c426541757 Now that 0 is overloaded, constant "zero" and its type class "zero" are
paulson
parents: 6921
diff changeset
    13
instance int :: add_monoid (IntDef.Zero_def,zadd_int0,zadd_int0_right)
a1c426541757 Now that 0 is overloaded, constant "zero" and its type class "zero" are
paulson
parents: 6921
diff changeset
    14
instance int :: add_group {|Auto_tac|}
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    15
instance int :: add_agroup (zadd_commute)
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    16
instance int :: ring (zmult_assoc,zadd_zmult_distrib2,zadd_zmult_distrib)
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    17
instance int :: cring (zmult_commute)
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    18
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    19
end