src/HOL/MicroJava/BV/Effect.thy
changeset 13006 51c5f3f11d16
parent 12974 7acfab8361d1
child 13066 b57d926d1de2
equal deleted inserted replaced
13005:42a54d6cec15 13006:51c5f3f11d16
    11 types
    11 types
    12   succ_type = "(p_count \<times> state_type option) list"
    12   succ_type = "(p_count \<times> state_type option) list"
    13 
    13 
    14 text {* Program counter of successor instructions: *}
    14 text {* Program counter of successor instructions: *}
    15 consts
    15 consts
    16   succs :: "instr => p_count => p_count list"
    16   succs :: "instr \<Rightarrow> p_count \<Rightarrow> p_count list"
    17 primrec 
    17 primrec 
    18   "succs (Load idx) pc         = [pc+1]"
    18   "succs (Load idx) pc         = [pc+1]"
    19   "succs (Store idx) pc        = [pc+1]"
    19   "succs (Store idx) pc        = [pc+1]"
    20   "succs (LitPush v) pc        = [pc+1]"
    20   "succs (LitPush v) pc        = [pc+1]"
    21   "succs (Getfield F C) pc     = [pc+1]"
    21   "succs (Getfield F C) pc     = [pc+1]"
    34   "succs (Invoke C mn fpTs) pc = [pc+1]"
    34   "succs (Invoke C mn fpTs) pc = [pc+1]"
    35   "succs Throw pc              = [pc]"
    35   "succs Throw pc              = [pc]"
    36 
    36 
    37 text "Effect of instruction on the state type:"
    37 text "Effect of instruction on the state type:"
    38 consts 
    38 consts 
    39 eff' :: "instr \<times> jvm_prog \<times> state_type => state_type"
    39 eff' :: "instr \<times> jvm_prog \<times> state_type \<Rightarrow> state_type"
    40 
    40 
    41 recdef eff' "{}"
    41 recdef eff' "{}"
    42 "eff' (Load idx,  G, (ST, LT))          = (ok_val (LT ! idx) # ST, LT)"
    42 "eff' (Load idx,  G, (ST, LT))          = (ok_val (LT ! idx) # ST, LT)"
    43 "eff' (Store idx, G, (ts#ST, LT))       = (ST, LT[idx:= OK ts])"
    43 "eff' (Store idx, G, (ts#ST, LT))       = (ST, LT[idx:= OK ts])"
    44 "eff' (LitPush v, G, (ST, LT))           = (the (typeof (\<lambda>v. None) v) # ST, LT)"
    44 "eff' (LitPush v, G, (ST, LT))           = (the (typeof (\<lambda>v. None) v) # ST, LT)"
    82 lemma match_some_entry:
    82 lemma match_some_entry:
    83   "match G X pc et = (if \<exists>e \<in> set et. match_exception_entry G X pc e then [X] else [])"
    83   "match G X pc et = (if \<exists>e \<in> set et. match_exception_entry G X pc e then [X] else [])"
    84   by (induct et) auto
    84   by (induct et) auto
    85 
    85 
    86 consts
    86 consts
    87   xcpt_names :: "instr \<times> jvm_prog \<times> p_count \<times> exception_table => cname list" 
    87   xcpt_names :: "instr \<times> jvm_prog \<times> p_count \<times> exception_table \<Rightarrow> cname list" 
    88 recdef xcpt_names "{}"
    88 recdef xcpt_names "{}"
    89   "xcpt_names (Getfield F C, G, pc, et) = match G (Xcpt NullPointer) pc et" 
    89   "xcpt_names (Getfield F C, G, pc, et) = match G (Xcpt NullPointer) pc et" 
    90   "xcpt_names (Putfield F C, G, pc, et) = match G (Xcpt NullPointer) pc et" 
    90   "xcpt_names (Putfield F C, G, pc, et) = match G (Xcpt NullPointer) pc et" 
    91   "xcpt_names (New C, G, pc, et)        = match G (Xcpt OutOfMemory) pc et"
    91   "xcpt_names (New C, G, pc, et)        = match G (Xcpt OutOfMemory) pc et"
    92   "xcpt_names (Checkcast C, G, pc, et)  = match G (Xcpt ClassCast) pc et"
    92   "xcpt_names (Checkcast C, G, pc, et)  = match G (Xcpt ClassCast) pc et"
   102        (xcpt_names (i,G,pc,et))"
   102        (xcpt_names (i,G,pc,et))"
   103 
   103 
   104   norm_eff :: "instr \<Rightarrow> jvm_prog \<Rightarrow> state_type option \<Rightarrow> state_type option"
   104   norm_eff :: "instr \<Rightarrow> jvm_prog \<Rightarrow> state_type option \<Rightarrow> state_type option"
   105   "norm_eff i G == option_map (\<lambda>s. eff' (i,G,s))"
   105   "norm_eff i G == option_map (\<lambda>s. eff' (i,G,s))"
   106 
   106 
   107   eff :: "instr => jvm_prog => p_count => exception_table => state_type option => succ_type"
   107   eff :: "instr \<Rightarrow> jvm_prog \<Rightarrow> p_count \<Rightarrow> exception_table \<Rightarrow> state_type option \<Rightarrow> succ_type"
   108   "eff i G pc et s == (map (\<lambda>pc'. (pc',norm_eff i G s)) (succs i pc)) @ (xcpt_eff i G pc s et)"
   108   "eff i G pc et s == (map (\<lambda>pc'. (pc',norm_eff i G s)) (succs i pc)) @ (xcpt_eff i G pc s et)"
   109 
   109 
   110 constdefs
   110 constdefs
   111   isPrimT :: "ty \<Rightarrow> bool"
   111   isPrimT :: "ty \<Rightarrow> bool"
   112   "isPrimT T == case T of PrimT T' \<Rightarrow> True | RefT T' \<Rightarrow> False"
   112   "isPrimT T == case T of PrimT T' \<Rightarrow> True | RefT T' \<Rightarrow> False"
   125   by (simp add: list_all2_def) 
   125   by (simp add: list_all2_def) 
   126 
   126 
   127 
   127 
   128 text "Conditions under which eff is applicable:"
   128 text "Conditions under which eff is applicable:"
   129 consts
   129 consts
   130 app' :: "instr \<times> jvm_prog \<times> p_count \<times> nat \<times> ty \<times> state_type => bool"
   130 app' :: "instr \<times> jvm_prog \<times> p_count \<times> nat \<times> ty \<times> state_type \<Rightarrow> bool"
   131 
   131 
   132 recdef app' "{}"
   132 recdef app' "{}"
   133 "app' (Load idx, G, pc, maxs, rT, s) = 
   133 "app' (Load idx, G, pc, maxs, rT, s) = 
   134   (idx < length (snd s) \<and> (snd s) ! idx \<noteq> Err \<and> length (fst s) < maxs)"
   134   (idx < length (snd s) \<and> (snd s) ! idx \<noteq> Err \<and> length (fst s) < maxs)"
   135 "app' (Store idx, G, pc, maxs, rT, (ts#ST, LT)) = 
   135 "app' (Store idx, G, pc, maxs, rT, (ts#ST, LT)) = 
   178 
   178 
   179 constdefs
   179 constdefs
   180   xcpt_app :: "instr \<Rightarrow> jvm_prog \<Rightarrow> nat \<Rightarrow> exception_table \<Rightarrow> bool"
   180   xcpt_app :: "instr \<Rightarrow> jvm_prog \<Rightarrow> nat \<Rightarrow> exception_table \<Rightarrow> bool"
   181   "xcpt_app i G pc et \<equiv> \<forall>C\<in>set(xcpt_names (i,G,pc,et)). is_class G C"
   181   "xcpt_app i G pc et \<equiv> \<forall>C\<in>set(xcpt_names (i,G,pc,et)). is_class G C"
   182 
   182 
   183   app :: "instr => jvm_prog => nat => ty => nat => exception_table => state_type option => bool"
   183   app :: "instr \<Rightarrow> jvm_prog \<Rightarrow> nat \<Rightarrow> ty \<Rightarrow> nat \<Rightarrow> exception_table \<Rightarrow> state_type option \<Rightarrow> bool"
   184   "app i G maxs rT pc et s == case s of None => True | Some t => app' (i,G,pc,maxs,rT,t) \<and> xcpt_app i G pc et"
   184   "app i G maxs rT pc et s == case s of None \<Rightarrow> True | Some t \<Rightarrow> app' (i,G,pc,maxs,rT,t) \<and> xcpt_app i G pc et"
   185 
   185 
   186 
   186 
   187 lemma 1: "2 < length a ==> (\<exists>l l' l'' ls. a = l#l'#l''#ls)"
   187 lemma 1: "2 < length a \<Longrightarrow> (\<exists>l l' l'' ls. a = l#l'#l''#ls)"
   188 proof (cases a)
   188 proof (cases a)
   189   fix x xs assume "a = x#xs" "2 < length a"
   189   fix x xs assume "a = x#xs" "2 < length a"
   190   thus ?thesis by - (cases xs, simp, cases "tl xs", auto)
   190   thus ?thesis by - (cases xs, simp, cases "tl xs", auto)
   191 qed auto
   191 qed auto
   192 
   192 
   193 lemma 2: "\<not>(2 < length a) ==> a = [] \<or> (\<exists> l. a = [l]) \<or> (\<exists> l l'. a = [l,l'])"
   193 lemma 2: "\<not>(2 < length a) \<Longrightarrow> a = [] \<or> (\<exists> l. a = [l]) \<or> (\<exists> l l'. a = [l,l'])"
   194 proof -;
   194 proof -;
   195   assume "\<not>(2 < length a)"
   195   assume "\<not>(2 < length a)"
   196   hence "length a < (Suc (Suc (Suc 0)))" by simp
   196   hence "length a < (Suc (Suc (Suc 0)))" by simp
   197   hence * : "length a = 0 \<or> length a = Suc 0 \<or> length a = Suc (Suc 0)" 
   197   hence * : "length a = 0 \<or> length a = Suc 0 \<or> length a = Suc (Suc 0)" 
   198     by (auto simp add: less_Suc_eq)
   198     by (auto simp add: less_Suc_eq)
   201     fix x 
   201     fix x 
   202     assume "length x = Suc 0"
   202     assume "length x = Suc 0"
   203     hence "\<exists> l. x = [l]"  by - (cases x, auto)
   203     hence "\<exists> l. x = [l]"  by - (cases x, auto)
   204   } note 0 = this
   204   } note 0 = this
   205 
   205 
   206   have "length a = Suc (Suc 0) ==> \<exists>l l'. a = [l,l']" by (cases a, auto dest: 0)
   206   have "length a = Suc (Suc 0) \<Longrightarrow> \<exists>l l'. a = [l,l']" by (cases a, auto dest: 0)
   207   with * show ?thesis by (auto dest: 0)
   207   with * show ?thesis by (auto dest: 0)
   208 qed
   208 qed
   209 
   209 
   210 lemmas [simp] = app_def xcpt_app_def
   210 lemmas [simp] = app_def xcpt_app_def
   211 
   211 
   335   method (G,C) (mn,fpTs) = Some (mD', rT', b') \<and> 
   335   method (G,C) (mn,fpTs) = Some (mD', rT', b') \<and> 
   336   (\<forall>C \<in> set (match_any G pc et). is_class G C))" (is "?app s = ?P s")
   336   (\<forall>C \<in> set (match_any G pc et). is_class G C))" (is "?app s = ?P s")
   337 proof (cases (open) s)
   337 proof (cases (open) s)
   338   note list_all2_def [simp]
   338   note list_all2_def [simp]
   339   case Pair
   339   case Pair
   340   have "?app (a,b) ==> ?P (a,b)"
   340   have "?app (a,b) \<Longrightarrow> ?P (a,b)"
   341   proof -
   341   proof -
   342     assume app: "?app (a,b)"
   342     assume app: "?app (a,b)"
   343     hence "a = (rev (rev (take (length fpTs) a))) @ (drop (length fpTs) a) \<and> 
   343     hence "a = (rev (rev (take (length fpTs) a))) @ (drop (length fpTs) a) \<and> 
   344            length fpTs < length a" (is "?a \<and> ?l") 
   344            length fpTs < length a" (is "?a \<and> ?l") 
   345       by (auto simp add: app_def)
   345       by (auto simp add: app_def)