diff -r eec2582923f6 -r b95d12325b51 src/HOL/MicroJava/J/Exceptions.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/MicroJava/J/Exceptions.thy Wed Oct 23 16:09:02 2002 +0200 @@ -0,0 +1,65 @@ +(* Title: HOL/MicroJava/J/Exceptions.thy + ID: $Id$ + Author: Gerwin Klein, Martin Strecker + Copyright 2002 Technische Universitaet Muenchen +*) + +theory Exceptions = State: + +text {* a new, blank object with default values in all fields: *} +constdefs + blank :: "'c prog \ cname \ obj" + "blank G C \ (C,init_vars (fields(G,C)))" + + start_heap :: "'c prog \ aheap" + "start_heap G \ empty (XcptRef NullPointer \ blank G (Xcpt NullPointer)) + (XcptRef ClassCast \ blank G (Xcpt ClassCast)) + (XcptRef OutOfMemory \ blank G (Xcpt OutOfMemory))" + + +consts + cname_of :: "aheap \ val \ cname" + +translations + "cname_of hp v" == "fst (the (hp (the_Addr v)))" + + +constdefs + preallocated :: "aheap \ bool" + "preallocated hp \ \x. \fs. hp (XcptRef x) = Some (Xcpt x, fs)" + +lemma preallocatedD: + "preallocated hp \ \fs. hp (XcptRef x) = Some (Xcpt x, fs)" + by (unfold preallocated_def) fast + +lemma preallocatedE [elim?]: + "preallocated hp \ (\fs. hp (XcptRef x) = Some (Xcpt x, fs) \ P hp) \ P hp" + by (fast dest: preallocatedD) + +lemma cname_of_xcp: + "raise_if b x None = Some xcp \ preallocated hp + \ cname_of (hp::aheap) xcp = Xcpt x" +proof - + assume "raise_if b x None = Some xcp" + hence "xcp = Addr (XcptRef x)" + by (simp add: raise_if_def split: split_if_asm) + moreover + assume "preallocated hp" + then obtain fs where "hp (XcptRef x) = Some (Xcpt x, fs)" .. + ultimately + show ?thesis by simp +qed + +lemma preallocated_start: + "preallocated (start_heap G)" + apply (unfold preallocated_def) + apply (unfold start_heap_def) + apply (rule allI) + apply (case_tac x) + apply (auto simp add: blank_def) + done + + + +end +