--- a/src/HOL/UNITY/Comp/Counterc.thy Tue Sep 19 23:15:32 2006 +0200
+++ b/src/HOL/UNITY/Comp/Counterc.thy Tue Sep 19 23:15:34 2006 +0200
@@ -16,7 +16,6 @@
theory Counterc imports "../UNITY_Main" begin
typedecl state
-arities state :: type
consts
C :: "state=>int"
--- a/src/HOL/UNITY/Comp/PriorityAux.thy Tue Sep 19 23:15:32 2006 +0200
+++ b/src/HOL/UNITY/Comp/PriorityAux.thy Tue Sep 19 23:15:34 2006 +0200
@@ -11,7 +11,6 @@
begin
typedecl vertex
-arities vertex :: type
constdefs
symcl :: "(vertex*vertex)set=>(vertex*vertex)set"