removed duplicate arities;
authorwenzelm
Tue, 19 Sep 2006 23:15:34 +0200
changeset 20625 1bb9a04f8c22
parent 20624 ba081ac0ed7e
child 20626 f5700d8ed070
removed duplicate arities;
src/HOL/UNITY/Comp/Counterc.thy
src/HOL/UNITY/Comp/PriorityAux.thy
--- 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"