# HG changeset patch # User paulson # Date 1114612784 -7200 # Node ID a344c4284972440dbdef8aad83eb3848d027ea51 # Parent 7bc8b9683224b349ff5ca2cc6e33f7e2ebac27ee partial modernising of theory headers diff -r 7bc8b9683224 -r a344c4284972 src/HOL/MicroJava/Comp/AuxLemmas.thy --- a/src/HOL/MicroJava/Comp/AuxLemmas.thy Wed Apr 27 11:49:20 2005 +0200 +++ b/src/HOL/MicroJava/Comp/AuxLemmas.thy Wed Apr 27 16:39:44 2005 +0200 @@ -6,7 +6,9 @@ (* Auxiliary Lemmas *) -theory AuxLemmas = JBasis: +theory AuxLemmas +imports "../J/JBasis" +begin (**********************************************************************) (* List.thy *) diff -r 7bc8b9683224 -r a344c4284972 src/HOL/MicroJava/Comp/CorrComp.thy --- 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) diff -r 7bc8b9683224 -r a344c4284972 src/HOL/MicroJava/JVM/JVMState.thy --- a/src/HOL/MicroJava/JVM/JVMState.thy Wed Apr 27 11:49:20 2005 +0200 +++ b/src/HOL/MicroJava/JVM/JVMState.thy Wed Apr 27 16:39:44 2005 +0200 @@ -9,7 +9,9 @@ \isaheader{State of the JVM} *} -theory JVMState = Conform: +theory JVMState +imports "../J/Conform" +begin section {* Frame Stack *} types