src/HOL/Bali/Decl.thy
changeset 18551 be0705186ff5
parent 18447 da548623916a
child 18576 8d98b7711e47
--- a/src/HOL/Bali/Decl.thy	Tue Jan 03 11:32:16 2006 +0100
+++ b/src/HOL/Bali/Decl.thy	Tue Jan 03 11:32:55 2006 +0100
@@ -324,7 +324,7 @@
          methods:: "mdecl list"
          init   :: "stmt"       --{* initializer *}
 
-record class = cbody +           --{* class *}
+record "class" = cbody +           --{* class *}
         super   :: "qtname"      --{* superclass *}
         superIfs:: "qtname list" --{* implemented interfaces *}
 types	
@@ -405,7 +405,7 @@
 
 syntax
   iface     :: "prog  \<Rightarrow> (qtname, iface) table"
-  class     :: "prog  \<Rightarrow> (qtname, class) table"
+  "class"     :: "prog  \<Rightarrow> (qtname, class) table"
   is_iface  :: "prog  \<Rightarrow> qtname  \<Rightarrow> bool"
   is_class  :: "prog  \<Rightarrow> qtname  \<Rightarrow> bool"