diff -r 73de0ef7cb25 -r 78f91fcdf560 src/HOL/MicroJava/BV/Effect.thy --- a/src/HOL/MicroJava/BV/Effect.thy Fri Nov 15 18:02:25 2002 +0100 +++ b/src/HOL/MicroJava/BV/Effect.thy Sat Nov 16 22:54:39 2002 +0100 @@ -73,23 +73,23 @@ if start_pc <= pc \ pc < end_pc then catch_type#es' else es')" consts - match :: "jvm_prog \ cname \ p_count \ exception_table \ cname list" + match :: "jvm_prog \ xcpt \ p_count \ exception_table \ cname list" primrec "match G X pc [] = []" "match G X pc (e#es) = - (if match_exception_entry G X pc e then [X] else match G X pc es)" + (if match_exception_entry G (Xcpt X) pc e then [Xcpt X] else match G X pc es)" lemma match_some_entry: - "match G X pc et = (if \e \ set et. match_exception_entry G X pc e then [X] else [])" + "match G X pc et = (if \e \ set et. match_exception_entry G (Xcpt X) pc e then [Xcpt X] else [])" by (induct et) auto consts xcpt_names :: "instr \ jvm_prog \ p_count \ exception_table \ cname list" recdef xcpt_names "{}" - "xcpt_names (Getfield F C, G, pc, et) = match G (Xcpt NullPointer) pc et" - "xcpt_names (Putfield F C, G, pc, et) = match G (Xcpt NullPointer) pc et" - "xcpt_names (New C, G, pc, et) = match G (Xcpt OutOfMemory) pc et" - "xcpt_names (Checkcast C, G, pc, et) = match G (Xcpt ClassCast) pc et" + "xcpt_names (Getfield F C, G, pc, et) = match G NullPointer pc et" + "xcpt_names (Putfield F C, G, pc, et) = match G NullPointer pc et" + "xcpt_names (New C, G, pc, et) = match G OutOfMemory pc et" + "xcpt_names (Checkcast C, G, pc, et) = match G ClassCast pc et" "xcpt_names (Throw, G, pc, et) = match_any G pc et" "xcpt_names (Invoke C m p, G, pc, et) = match_any G pc et" "xcpt_names (i, G, pc, et) = []" @@ -257,26 +257,26 @@ lemma appGetField[simp]: "(app (Getfield F C) G maxs rT pc et (Some s)) = (\ oT vT ST LT. s = (oT#ST, LT) \ is_class G C \ - field (G,C) F = Some (C,vT) \ G \ oT \ (Class C) \ (\x \ set (match G (Xcpt NullPointer) pc et). is_class G x))" + field (G,C) F = Some (C,vT) \ G \ oT \ (Class C) \ (\x \ set (match G NullPointer pc et). is_class G x))" by (cases s, cases "2 vT vT' oT ST LT. s = (vT#oT#ST, LT) \ is_class G C \ field (G,C) F = Some (C, vT') \ G \ oT \ (Class C) \ G \ vT \ vT' \ - (\x \ set (match G (Xcpt NullPointer) pc et). is_class G x))" + (\x \ set (match G NullPointer pc et). is_class G x))" by (cases s, cases "2 ST LT. s=(ST,LT) \ is_class G C \ length ST < maxs \ - (\x \ set (match G (Xcpt OutOfMemory) pc et). is_class G x))" + (\x \ set (match G OutOfMemory pc et). is_class G x))" by (cases s, simp) lemma appCheckcast[simp]: "(app (Checkcast C) G maxs rT pc et (Some s)) = (\rT ST LT. s = (RefT rT#ST,LT) \ is_class G C \ - (\x \ set (match G (Xcpt ClassCast) pc et). is_class G x))" + (\x \ set (match G ClassCast pc et). is_class G x))" by (cases s, cases "fst s", simp add: app_def) (cases "hd (fst s)", auto) lemma appPop[simp]: