Up to index of Isabelle/HOL/objinit
theory EffectMono = Effect:(* Title: HOL/MicroJava/BV/EffMono.thy ID: $Id: EffectMono.html,v 1.1 2002/11/28 14:12:09 kleing Exp $ Author: Gerwin Klein Copyright 2000 Technische Universitaet Muenchen *) header {* \isaheader{Monotonicity of eff and app} *} theory EffectMono = Effect: lemma PrimT_PrimT: "(G \<turnstile> xb \<preceq> PrimT p) = (xb = PrimT p)" by (auto elim: widen.elims) lemma InitPrimT_InitPrimT: "(G \<turnstile> xb \<preceq>i Init (PrimT p)) = (xb = Init (PrimT p))" by (cases xb, auto elim: widen.elims simp add: subtype_def) lemma sup_loc_some [rule_format]: "\<forall>y n. (G \<turnstile> b <=l y) \<longrightarrow> n < length y \<longrightarrow> y!n = OK t \<longrightarrow> (\<exists>t. b!n = OK t \<and> (G \<turnstile> (b!n) <=o (y!n)))" (is "?P b") proof (induct (open) ?P b) show "?P []" by simp case Cons show "?P (a#list)" proof (clarsimp simp add: list_all2_Cons1 sup_loc_def Listn.le_def lesub_def) fix z zs n assume * : "G \<turnstile> a <=o z" "list_all2 (sup_ty_opt G) list zs" "n < Suc (length list)" "(z # zs) ! n = OK t" show "(\<exists>t. (a # list) ! n = OK t) \<and> G \<turnstile>(a # list) ! n <=o OK t" proof (cases n) case 0 with * show ?thesis by (clarsimp simp add: sup_ty_opt_OK) next case Suc with Cons * show ?thesis by (simp add: sup_loc_def Listn.le_def lesub_def) qed qed qed lemma all_set_conv_sup_loc [rule_format]: "\<forall>b. length a = length b \<longrightarrow> (\<forall>(x,y)\<in>set (zip a b). G \<turnstile> x \<preceq>i Init y) = (G \<turnstile> (map OK a) <=l (map OK (map Init b)))" (is "\<forall>b. length a = length b \<longrightarrow> ?Q a b" is "?P a") proof (induct "a") show "?P []" by simp fix l ls assume Cons: "?P ls" show "?P (l#ls)" proof (intro allI impI) fix b assume "length (l # ls) = length (b::ty list)" with Cons show "?Q (l # ls) b" by - (cases b, auto) qed qed lemma replace_UnInit: "\<lbrakk> G \<turnstile> a <=l b; X = UnInit C pc \<or> X = PartInit D \<rbrakk> \<Longrightarrow> G \<turnstile> (replace (OK X) v a) <=l (replace (OK X) v b)" proof - assume X: "X = UnInit C pc \<or> X = PartInit D" assume "G \<turnstile> a <=l b" then obtain l: "length a = length b" and a: "\<forall>i. i < length a \<longrightarrow> G \<turnstile> a!i <=o b!i" by (unfold sup_loc_def Listn.le_def lesub_def) (auto simp add: list_all2_conv_all_nth) { fix i assume "i < length (replace (OK X) v a)" hence i: "i < length a" by (simp add: replace_def) hence G: "G \<turnstile> a!i <=o b!i" by (simp add: a) from l i have i2: "i < length b" by simp have "G \<turnstile> (replace (OK X) v a)!i <=o (replace (OK X) v b)!i" proof (cases "a!i = OK X") case True with G i i2 X have "b!i = OK X \<or> b!i = Err" by (simp add: sup_ty_opt_def Err.le_def lesub_def init_le_def split: err.splits init_ty.split_asm) with True i i2 show ?thesis by (auto simp add: replace_def) (simp add: sup_ty_opt_def Err.le_def) next case False with G i i2 X have "b!i \<noteq> OK X" by (auto simp add: sup_ty_opt_def Err.le_def lesub_def init_le_def split: init_ty.split_asm err.splits) with False i i2 G show ?thesis by (simp add: replace_def) qed } with l show ?thesis by (unfold sup_loc_def Listn.le_def lesub_def) (auto simp add: list_all2_conv_all_nth replace_def) qed (* fixme: merge with above *) lemma replace_mapOK_UnInit: "\<lbrakk> G \<turnstile> map OK ST <=l map OK ST'; X = UnInit C pc \<or> X = PartInit D \<rbrakk> \<Longrightarrow> G \<turnstile> map OK (replace X v ST) <=l map OK (replace X v ST')" proof - assume X: "X = UnInit C pc \<or> X = PartInit D" assume "G \<turnstile> map OK ST <=l map OK ST'" then obtain l: "length ST = length ST'" and a: "\<forall>i. i < length ST \<longrightarrow> G \<turnstile> ST!i \<preceq>i ST'!i" by (unfold sup_loc_def Listn.le_def lesub_def) (auto simp add: list_all2_conv_all_nth) { fix i assume "i < length (replace X v ST)" hence i: "i < length ST" by (simp add: replace_def) hence G: "G \<turnstile> ST!i \<preceq>i ST'!i" by (simp add: a) from l i have i2: "i < length ST'" by simp have "G \<turnstile> (replace X v ST)!i \<preceq>i (replace X v ST')!i" proof (cases "ST!i = X") case True with G i i2 X have "ST'!i = X" by (simp add: init_le_def split: init_ty.split_asm) with True i i2 show ?thesis by (simp add: replace_def) next case False with G i i2 X have "ST'!i \<noteq> X" by (auto simp add: init_le_def split: init_ty.split_asm) with False i i2 G show ?thesis by (simp add: replace_def) qed } with l show ?thesis by (unfold sup_loc_def Listn.le_def lesub_def) (auto simp add: list_all2_conv_all_nth replace_def) qed lemma append_length_n [rule_format]: "\<forall>n. n \<le> length x \<longrightarrow> (\<exists>a b. x = a@b \<and> length a = n)" (is "?P x") proof (induct (open) ?P x) show "?P []" by simp fix l ls assume Cons: "?P ls" show "?P (l#ls)" proof (intro allI impI) fix n assume l: "n \<le> length (l # ls)" show "\<exists>a b. l # ls = a @ b \<and> length a = n" proof (cases n) assume "n=0" thus ?thesis by simp next fix "n'" assume s: "n = Suc n'" with l have "n' \<le> length ls" by simp hence "\<exists>a b. ls = a @ b \<and> length a = n'" by (rule Cons [rule_format]) thus ?thesis proof (elim exE conjE) fix a b assume "ls = a @ b" "length a = n'" with s have "l # ls = (l#a) @ b \<and> length (l#a) = n" by simp thus ?thesis by blast qed qed qed qed lemma rev_append_cons: "n < length x \<Longrightarrow> \<exists>a b c. x = (rev a) @ b # c \<and> length a = n" proof - assume n: "n < length x" hence "n \<le> length x" by simp hence "\<exists>a b. x = a @ b \<and> length a = n" by (rule append_length_n) thus ?thesis proof (elim exE conjE) fix r d assume x: "x = r@d" "length r = n" with n have "\<exists>b c. d = b#c" by (simp add: neq_Nil_conv) thus ?thesis proof (elim exE conjE) fix b c assume "d = b#c" with x have "x = (rev (rev r)) @ b # c \<and> length (rev r) = n" by simp thus ?thesis by blast qed qed qed lemmas [iff] = not_Err_eq lemma UnInit_set: "\<lbrakk> G \<turnstile> map OK ST <=l map OK ST'; UnInit C pc \<notin> set ST' \<rbrakk> \<Longrightarrow> UnInit C pc \<notin> set ST" proof assume "UnInit C pc \<in> set ST" then obtain x y where "ST = x @ UnInit C pc # y" by (clarsimp simp add: in_set_conv_decomp) hence l: "length x < length (map OK ST) \<and> ST!length x = UnInit C pc" by (simp add: nth_append) moreover assume G: "G \<turnstile> map OK ST <=l map OK ST'" moreover from G have lm: "length (map OK ST) = length (map OK ST')" by (rule sup_loc_length) ultimately obtain T where "UnInit C pc = ST'!length x" by clarify (drule sup_loc_nth, assumption, clarsimp simp add: nth_map) with G l lm [symmetric] have "UnInit C pc \<in> set ST'" by (auto simp add: set_conv_nth sup_loc_length) moreover assume "UnInit C pc \<notin> set ST'" ultimately show False by blast qed lemma sup_loc_length_map: "G \<turnstile> map f a <=l map g b \<Longrightarrow> length a = length b" proof - assume "G \<turnstile> map f a <=l map g b" hence "length (map f a) = length (map g b)" by (rule sup_loc_length) thus ?thesis by simp qed lemma app_mono: "\<lbrakk>G \<turnstile> s <=' s'; app i G C pc m rT ini et s'\<rbrakk> \<Longrightarrow> app i G C pc m rT ini et s" proof - { fix s1 s2 z assume G: "G \<turnstile> s2 <=s s1" assume app: "app i G C pc m rT ini et (Some (s1,z))" note [simp] = sup_loc_length sup_loc_length_map have "app i G C pc m rT ini et (Some (s2,z))" proof (cases (open) i) case Load from G Load app have "G \<turnstile> snd s2 <=l snd s1" by (auto simp add: sup_state_conv) with G Load app show ?thesis by (cases s2) (auto simp add: sup_state_conv dest: sup_loc_some) next case Store with G app show ?thesis by (cases s2, auto simp add: map_eq_Cons sup_loc_Cons2 sup_state_conv) next case LitPush with G app show ?thesis by (cases s2, auto simp add: sup_state_conv) next case New with G app show ?thesis by (cases s2, auto simp add: sup_state_conv dest: UnInit_set) next case Getfield with app G show ?thesis by (cases s2, clarsimp simp add: sup_state_Cons2) (rule init_trans) next case Putfield with app obtain vT oT ST LT b where s1: "s1 = (vT # oT # ST, LT)" and "field (G, cname) vname = Some (cname, b)" "is_class G cname" and oT: "G\<turnstile> oT \<preceq>i (Init (Class cname))" and vT: "G\<turnstile> vT \<preceq>i (Init b)" and xc: "Ball (set (match G (Xcpt NullPointer) pc et)) (is_class G)" by force moreover from s1 G obtain vT' oT' ST' LT' where s2: "s2 = (vT' # oT' # ST', LT')" and oT': "G\<turnstile> oT' \<preceq>i oT" and vT': "G\<turnstile> vT' \<preceq>i vT" by (cases s2, auto simp add: sup_state_Cons2) moreover from vT' vT have "G \<turnstile> vT' \<preceq>i (Init b)" by (rule init_trans) moreover from oT' oT have "G\<turnstile> oT' \<preceq>i (Init (Class cname))" by (rule init_trans) ultimately show ?thesis by (cases s2, auto simp add: Putfield xc) next case Checkcast with app G show ?thesis by (cases s2, auto simp add: init_le_Init2 sup_state_Cons2 intro!: widen_RefT2) next case Return with app G show ?thesis by (cases s2, auto simp add: sup_state_Cons2) (rule init_trans, assumption+, rule init_trans) next case Pop with app G show ?thesis by (cases s2, clarsimp simp add: sup_state_Cons2) next case Dup with app G show ?thesis by (cases s2, clarsimp simp add: sup_state_Cons2, auto dest: sup_state_length) next case Dup_x1 with app G show ?thesis by (cases s2, clarsimp simp add: sup_state_Cons2, auto dest: sup_state_length) next case Dup_x2 with app G show ?thesis by (cases s2, clarsimp simp add: sup_state_Cons2, auto dest: sup_state_length) next case Swap with app G show ?thesis by (cases s2, clarsimp simp add: sup_state_Cons2) next case IAdd with app G show ?thesis by (cases s2, auto simp add: sup_state_Cons2 InitPrimT_InitPrimT) next case Goto with app show ?thesis by simp next case Ifcmpeq with app G show ?thesis apply (cases s2, auto simp add: sup_state_Cons2 InitPrimT_InitPrimT) apply (auto simp add: init_le_Init2 widen_RefT2) done next case Invoke with app obtain apTs X ST LT mD' rT' b' where s1: "s1 = (rev apTs @ X # ST, LT)" and l: "length apTs = length list" and c: "is_class G cname" and w: "\<forall>(x,y) \<in> set (zip apTs list). G \<turnstile> x \<preceq>i Init y" and m: "method (G, cname) (mname, list) = Some (mD', rT', b')" and mn: "mname \<noteq> init" and C: "G \<turnstile> X \<preceq>i Init (Class cname)" and x: "\<forall>C \<in> set (match_any G pc et). is_class G C" by simp blast obtain apTs' X' ST' LT' where s2: "s2 = (rev apTs' @ X' # ST', LT')" and l': "length apTs' = length list" proof - from l s1 G have "length list < length (fst s2)" by simp hence "\<exists>a b c. (fst s2) = rev a @ b # c \<and> length a = length list" by (rule rev_append_cons [rule_format]) thus ?thesis by (cases s2, elim exE conjE, simp) (rule that) qed from l l' have "length (rev apTs') = length (rev apTs)" by simp from this s1 s2 G obtain G': "G \<turnstile> (apTs',LT') <=s (apTs,LT)" and X : "G \<turnstile> X' \<preceq>i X" and "G \<turnstile> (ST',LT') <=s (ST,LT)" by (simp add: sup_state_rev_fst sup_state_append_fst sup_state_Cons1) with C have C': "G \<turnstile> X' \<preceq>i Init (Class cname)" by (blast intro: init_trans) from G' have "G \<turnstile> map OK apTs' <=l map OK apTs" by (simp add: sup_state_conv) also from l w have "G \<turnstile> map OK apTs <=l map OK (map Init list)" by (simp add: all_set_conv_sup_loc) finally have "G \<turnstile> map OK apTs' <=l map OK (map Init list)" . with l' have w': "\<forall>(x,y) \<in> set (zip apTs' list). G \<turnstile> x \<preceq>i Init y" by (simp add: all_set_conv_sup_loc) from Invoke s2 l' w' C' m c mn x show ?thesis by (simp del: split_paired_Ex) blast next case Invoke_special with app obtain apTs X ST LT rT' b' where s1: "s1 = (rev apTs @ X # ST, LT)" and l: "length apTs = length list" and c: "is_class G cname" and w: "\<forall>(x,y) \<in> set (zip apTs list). G \<turnstile> x \<preceq>i Init y" and m: "method (G, cname) (mname, list) = Some (cname, rT', b')" and mn: "mname = init" and C: "(\<exists>pc. X = UnInit cname pc) \<or> (X = PartInit C \<and> G \<turnstile> C \<prec>C1 cname \<and> ¬ z)" and x: "\<forall>C \<in> set (match_any G pc et). is_class G C" by simp blast obtain apTs' X' ST' LT' where s2: "s2 = (rev apTs' @ X' # ST', LT')" and l': "length apTs' = length list" proof - from l s1 G have "length list < length (fst s2)" by simp hence "\<exists>a b c. (fst s2) = rev a @ b # c \<and> length a = length list" by (rule rev_append_cons [rule_format]) thus ?thesis by (cases s2, elim exE conjE, simp) (rule that) qed from l l' have "length (rev apTs') = length (rev apTs)" by simp from this s1 s2 G obtain G': "G \<turnstile> (apTs',LT') <=s (apTs,LT)" and X : "G \<turnstile> X' \<preceq>i X" and "G \<turnstile> (ST',LT') <=s (ST,LT)" by (simp add: sup_state_rev_fst sup_state_append_fst sup_state_Cons1) with C have C': "(\<exists>pc. X' = UnInit cname pc) \<or> (X' = PartInit C \<and> G \<turnstile> C \<prec>C1 cname \<and> ¬ z)" by auto from G' have "G \<turnstile> map OK apTs' <=l map OK apTs" by (simp add: sup_state_conv) also from l w have "G \<turnstile> map OK apTs <=l map OK (map Init list)" by (simp add: all_set_conv_sup_loc) finally have "G \<turnstile> map OK apTs' <=l map OK (map Init list)" . with l' have w': "\<forall>(x,y) \<in> set (zip apTs' list). G \<turnstile> x \<preceq>i Init y" by (simp add: all_set_conv_sup_loc) from Invoke_special s2 l' w' C' m c mn x show ?thesis by (simp del: split_paired_Ex) blast next case Throw with app G show ?thesis by (cases s2, clarsimp simp add: sup_state_Cons2 init_le_Init2 widen_RefT2) qed } note this [simp] assume "G \<turnstile> s <=' s'" "app i G C pc m rT ini et s'" thus ?thesis by (cases s, cases s', auto) qed lemmas [simp del] = split_paired_Ex lemma eff_bool_mono: "\<lbrakk> app i G C pc m rT ini et (Some t); G \<turnstile> s <=b t \<rbrakk> \<Longrightarrow> G \<turnstile> eff_bool i G pc s <=b eff_bool i G pc t" proof - obtain s1 z where s': "s = (s1,z)" by (cases s) moreover assume "G \<turnstile> s <=b t" ultimately obtain s2 where t: "t = (s2,z)" and G: "G \<turnstile> s1 <=s s2" by (cases t, simp) from s' t obtain a1 b1 a2 b2 where s12: "s1 = (a1,b1)" "s2 = (a2,b2)" by (cases s1, cases s2) moreover assume "app i G C pc m rT ini et (Some t)" moreover note t ultimately have app2: "app i G C pc m rT ini et (Some (s2,z))" by simp have "G \<turnstile> (Some (s1,z)) <=' (Some (s2,z))" by simp from this app2 have app1: "app i G C pc m rT ini et (Some (s1,z))" by (rule app_mono) note [simp] = eff_def eff_bool_def note s = s' t s12 show ?thesis proof (cases (open) i) case Load with s app1 obtain y where y: "nat < length b1" "b1 ! nat = OK y" by clarsimp from Load s app2 obtain y' where y': "nat < length b2" "b2 ! nat = OK y'" by clarsimp from G s have "G \<turnstile> b1 <=l b2" by (simp add: sup_state_conv) with y y' have "G \<turnstile> y \<preceq>i y'" by - (drule sup_loc_some, simp+) with Load G y y' s app1 app2 show ?thesis by (clarsimp simp add: sup_state_conv) next case Store with G s app1 app2 show ?thesis by (clarsimp simp add: sup_state_conv sup_loc_update) next case LitPush with G s app1 app2 show ?thesis by (clarsimp simp add: sup_state_Cons1) next case New with G s app1 app2 show ?thesis by (clarsimp simp add: sup_state_Cons1 sup_state_conv) (blast intro: replace_UnInit) next case Getfield with G s app1 app2 show ?thesis by (clarsimp simp add: sup_state_Cons1) next case Putfield with G s app1 app2 show ?thesis by (clarsimp simp add: sup_state_Cons1) next case Checkcast with G s app1 app2 show ?thesis by (clarsimp simp add: sup_state_Cons1) next case Invoke with s app1 obtain a X ST where s1: "s1 = (rev a @ X # ST, b1)" and l: "length a = length list" and C: "G \<turnstile> X \<preceq>i Init (Class cname)" by (clarsimp, blast) from Invoke s app2 obtain a' X' ST' where s2: "s2 = (rev a' @ X' # ST', b2)" and l': "length a' = length list" and C': "G \<turnstile> X' \<preceq>i Init (Class cname)" by (clarsimp, blast) from l l' have "length a = length a'" by simp from this G s1 s2 have "G \<turnstile> (ST, b1) <=s (ST', b2)" by (simp add: sup_state_append_fst sup_state_Cons1) with Invoke G app1 app2 s s1 s2 l l' show ?thesis by (clarsimp simp add: sup_state_conv) next case Return with G s show ?thesis by simp next case Pop with G s app1 app2 show ?thesis by (clarsimp simp add: sup_state_Cons1) next case Dup with G s app1 app2 show ?thesis by (clarsimp simp add: sup_state_Cons1) next case Dup_x1 with G s app1 app2 show ?thesis by (clarsimp simp add: sup_state_Cons1) next case Dup_x2 with G s app1 app2 show ?thesis by (clarsimp simp add: sup_state_Cons1) next case Swap with G s app1 app2 show ?thesis by (clarsimp simp add: sup_state_Cons1) next case IAdd with G s app1 app2 show ?thesis by (clarsimp simp add: sup_state_Cons1) next case Goto with G s app1 app2 show ?thesis by simp next case Ifcmpeq with G s app1 app2 show ?thesis by (clarsimp simp add: sup_state_Cons1) next case Invoke_special with s app1 obtain a X ST where s1: "s1 = (rev a @ X # ST, b1)" and l: "length a = length list" and C: "(\<exists>pc. X = UnInit cname pc) \<or> X = PartInit C \<and> G \<turnstile> C \<prec>C1 cname \<and> ¬ z" by (clarsimp, blast) from Invoke_special s app2 obtain a' X' ST' where s2: "s2 = (rev a' @ X' # ST', b2)" and l': "length a' = length list" and C': "(\<exists>pc. X' = UnInit cname pc) \<or> X' = PartInit C \<and> G \<turnstile> C \<prec>C1 cname \<and> ¬ z" by (clarsimp, blast) from l l' have lr: "length a = length a'" by simp from lr G s1 s2 have "G \<turnstile> (ST, b1) <=s (ST', b2)" by (simp add: sup_state_append_fst sup_state_Cons1) moreover from lr G s1 s2 have "G \<turnstile> X \<preceq>i X'" by (simp add: sup_state_append_fst sup_state_Cons1) with C C' have XX': "X = X'" by auto moreover note Invoke_special G app1 app2 s s1 s2 l l' C C' ultimately show ?thesis by (clarsimp simp add: sup_state_conv nth_append) (blast intro: replace_UnInit replace_mapOK_UnInit) next case Throw with G s show ?thesis by simp qed qed lemmas [iff del] = not_Err_eq end
lemma PrimT_PrimT:
G |- xb <= PrimT p = (xb = PrimT p)
lemma InitPrimT_InitPrimT:
G |- xb <=i Init (PrimT p) = (xb = Init (PrimT p))
lemma sup_loc_some:
[| G |- b <=l y; n < length y; y ! n = OK t |] ==> EX t. b ! n = OK t & G |- b ! n <=o y ! n
lemma all_set_conv_sup_loc:
length a = length b ==> (ALL (x, y):set (zip a b). G |- x <=i Init y) = G |- map OK a <=l map OK (map Init b)
lemma replace_UnInit:
[| G |- a <=l b; X = UnInit C pc | X = PartInit D |] ==> G |- replace (OK X) v a <=l replace (OK X) v b
lemma replace_mapOK_UnInit:
[| G |- map OK ST <=l map OK ST'; X = UnInit C pc | X = PartInit D |] ==> G |- map OK (replace X v ST) <=l map OK (replace X v ST')
lemma append_length_n:
n <= length x ==> EX a b. x = a @ b & length a = n
lemma rev_append_cons:
n < length x ==> EX a b c. x = rev a @ b # c & length a = n
lemmas
(x ~= Err) = (EX a. x = OK a)
lemma UnInit_set:
[| G |- map OK ST <=l map OK ST'; UnInit C pc ~: set ST' |] ==> UnInit C pc ~: set ST
lemma sup_loc_length_map:
G |- map f a <=l map g b ==> length a = length b
lemma app_mono:
[| G |- s <=' s'; app i G C pc m rT ini et s' |] ==> app i G C pc m rT ini et s
lemmas
(EX x. P x) = (EX a b. P (a, b))
lemma eff_bool_mono:
[| app i G C pc m rT ini et (Some t); G |- s <=b t |] ==> G |- eff_bool i G pc s <=b eff_bool i G pc t
lemmas
(x ~= Err) = (EX a. x = OK a)