src/HOL/MicroJava/J/Example.thy
changeset 11701 3d51fbf81c17
parent 11643 0b3a02daf7fb
child 11908 82f68fd05094
--- a/src/HOL/MicroJava/J/Example.thy	Fri Oct 05 21:50:37 2001 +0200
+++ b/src/HOL/MicroJava/J/Example.thy	Fri Oct 05 21:52:39 2001 +0200
@@ -100,7 +100,7 @@
 			     [(vee, PrimT Boolean)], 
 			     [((foo,[Class Base]),Class Base,foo_Base)]))"
   foo_Ext_def:"foo_Ext == ([x],[],Expr( {Ext}Cast Ext
-				       (LAcc x)..vee:=Lit (Intg #1)),
+				       (LAcc x)..vee:=Lit (Intg Numeral1)),
 				   Lit Null)"
   ExtC_def: "ExtC  == (Ext,  (Base  , 
 			     [(vee, PrimT Integer)], 
@@ -127,7 +127,7 @@
   "NP"   == "NullPointer"
   "tprg" == "[ObjectC, BaseC, ExtC]"
   "obj1"    <= "(Ext, empty((vee, Base)\<mapsto>Bool False)
-			   ((vee, Ext )\<mapsto>Intg #0))"
+			   ((vee, Ext )\<mapsto>Intg Numeral0))"
   "s0" == " Norm    (empty, empty)"
   "s1" == " Norm    (empty(a\<mapsto>obj1),empty(e\<mapsto>Addr a))"
   "s2" == " Norm    (empty(a\<mapsto>obj1),empty(x\<mapsto>Null)(This\<mapsto>Addr a))"