src/HOL/MicroJava/Comp/CorrComp.thy
changeset 15860 a344c4284972
parent 15481 fc075ae929e4
child 15866 f011452b2815
--- a/src/HOL/MicroJava/Comp/CorrComp.thy	Wed Apr 27 11:49:20 2005 +0200
+++ b/src/HOL/MicroJava/Comp/CorrComp.thy	Wed Apr 27 16:39:44 2005 +0200
@@ -5,7 +5,9 @@
 
 (* Compiler correctness statement and proof *)
 
-theory CorrComp =  JTypeSafe + LemmasComp:
+theory CorrComp
+imports "../J/JTypeSafe" "../Comp/LemmasComp"
+begin
 
 declare wf_prog_ws_prog [simp add]
 
@@ -1146,7 +1148,11 @@
 apply (subgoal_tac "(pns, lvars, blk, res) = gmb G md (mn, pTs)")
 apply (simp (no_asm_simp))
 apply (simp only: gh_conv)
-apply ((drule mp, rule TrueI)+, (drule spec)+, (drule mp, assumption)+, assumption)
+apply ((drule mp [OF _ TrueI])+, (drule spec)+)
+apply ((drule mp, assumption)+, assumption)
+  --{*extremely slow with Poly/ML (72s) and (under some circumstances)
+      unbearably slow with SML/NJ (up to 83 minutes)*}
+
 apply (simp (no_asm_use))
 apply (simp (no_asm_simp) add: gmb_def)