src/HOL/BCV/Types.thy
author paulson
Fri, 05 Nov 1999 12:45:37 +0100
changeset 7999 7acf6eb8eec1
parent 7626 5997f35954d7
permissions -rw-r--r--
Algebra and Polynomial theories, by Clemens Ballarin
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
     1
(*  Title:      HOL/BCV/Types.thy
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
     2
    ID:         $Id$
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
     3
    Author:     Tobias Nipkow
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
     4
    Copyright   1999 TUM
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
     5
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
     6
Types for the register machine
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
     7
*)
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
     8
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
     9
Types = Types0 +
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    10
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    11
instance typ :: order (le_typ_refl,le_typ_trans,le_typ_antisym,less_le_typ)
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    12
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    13
end