diff -r 4334c91b7405 -r f682f3f7b726 src/HOL/IMP/Abs_Int1.thy --- a/src/HOL/IMP/Abs_Int1.thy Wed Nov 23 23:31:32 2011 +0100 +++ b/src/HOL/IMP/Abs_Int1.thy Thu Nov 24 19:58:37 2011 +0100 @@ -1,7 +1,7 @@ (* Author: Tobias Nipkow *) theory Abs_Int1 -imports Abs_Int0_const +imports Abs_Int0 begin instantiation prod :: (preord,preord) preord @@ -37,20 +37,19 @@ end - locale Val_abs1_rep = Val_abs rep num' plus' - for rep :: "'a::L_top_bot \ val set" and num' plus' + + for rep :: "'a::L_top_bot \ val set" ("\") and num' plus' + assumes inter_rep_subset_rep_meet: "rep a1 \ rep a2 \ rep(a1 \ a2)" -and rep_Bot: "rep \ = {}" +and rep_Bot[simp]: "rep \ = {}" begin -lemma in_rep_meet: "x <: a1 \ x <: a2 \ x <: a1 \ a2" +lemma in_rep_meet: "x : \ a1 \ x : \ a2 \ x : \(a1 \ a2)" by (metis IntI inter_rep_subset_rep_meet set_mp) -lemma rep_meet[simp]: "rep(a1 \ a2) = rep a1 \ rep a2" -by (metis equalityI inter_rep_subset_rep_meet le_inf_iff le_rep meet_le1 meet_le2) +lemma rep_meet[simp]: "\(a1 \ a2) = \ a1 \ \ a2" +by (metis equalityI inter_rep_subset_rep_meet le_inf_iff mono_rep meet_le1 meet_le2) end @@ -59,76 +58,74 @@ fixes filter_plus' :: "'a \ 'a \ 'a \ 'a * 'a" and filter_less' :: "bool \ 'a \ 'a \ 'a * 'a" assumes filter_plus': "filter_plus' a a1 a2 = (a1',a2') \ - n1 <: a1 \ n2 <: a2 \ n1+n2 <: a \ n1 <: a1' \ n2 <: a2'" + n1 : \ a1 \ n2 : \ a2 \ n1+n2 : \ a \ n1 : \ a1' \ n2 : \ a2'" and filter_less': "filter_less' (n1 - n1 <: a1 \ n2 <: a2 \ n1 <: a1' \ n2 <: a2'" + n1 : \ a1 \ n2 : \ a2 \ n1 : \ a1' \ n2 : \ a2'" locale Abs_Int1 = Val_abs1 begin -lemma in_rep_join_UpI: "s <:up S1 | s <:up S2 \ s <:up S1 \ S2" -by (metis join_ge1 join_ge2 up_fun_in_rep_le) +lemma in_rep_join_UpI: "s : \\<^isub>u S1 | s : \\<^isub>u S2 \ s : \\<^isub>u(S1 \ S2)" +by (metis (no_types) join_ge1 join_ge2 mono_rep_u set_rev_mp) -fun aval' :: "aexp \ 'a st up \ 'a" where -"aval' _ Bot = \" | -"aval' (N n) _ = num' n" | -"aval' (V x) (Up S) = lookup S x" | -"aval' (Plus a1 a2) S = plus' (aval' a1 S) (aval' a2 S)" +fun aval'' :: "aexp \ 'a st option \ 'a" where +"aval'' e None = \" | +"aval'' e (Some sa) = aval' e sa" -lemma aval'_sound: "s <:up S \ aval a s <: aval' a S" -by(induct a)(auto simp: rep_num' rep_plus' in_rep_up_iff lookup_def rep_st_def) +lemma aval''_sound: "s : \\<^isub>u S \ aval a s : \(aval'' a S)" +by(cases S)(simp add: aval'_sound)+ subsubsection "Backward analysis" -fun afilter :: "aexp \ 'a \ 'a st up \ 'a st up" where -"afilter (N n) a S = (if n <: a then S else Bot)" | -"afilter (V x) a S = (case S of Bot \ Bot | Up S \ +fun afilter :: "aexp \ 'a \ 'a st option \ 'a st option" where +"afilter (N n) a S = (if n : \ a then S else None)" | +"afilter (V x) a S = (case S of None \ None | Some S \ let a' = lookup S x \ a in - if a' \ \ then Bot else Up(update S x a'))" | + if a' \ \ then None else Some(update S x a'))" | "afilter (Plus e1 e2) a S = - (let (a1,a2) = filter_plus' a (aval' e1 S) (aval' e2 S) + (let (a1,a2) = filter_plus' a (aval'' e1 S) (aval'' e2 S) in afilter e1 a1 (afilter e2 a2 S))" -text{* The test for @{const Bot} in the @{const V}-case is important: @{const -Bot} indicates that a variable has no possible values, i.e.\ that the current +text{* The test for @{const bot} in the @{const V}-case is important: @{const +bot} indicates that a variable has no possible values, i.e.\ that the current program point is unreachable. But then the abstract state should collapse to -@{const bot}. Put differently, we maintain the invariant that in an abstract -state all variables are mapped to non-@{const Bot} values. Otherwise the -(pointwise) join of two abstract states, one of which contains @{const Bot} -values, may produce too large a result, thus making the analysis less -precise. *} +@{const None}. Put differently, we maintain the invariant that in an abstract +state of the form @{term"Some s"}, all variables are mapped to non-@{const +bot} values. Otherwise the (pointwise) join of two abstract states, one of +which contains @{const bot} values, may produce too large a result, thus +making the analysis less precise. *} -fun bfilter :: "bexp \ bool \ 'a st up \ 'a st up" where -"bfilter (Bc v) res S = (if v=res then S else Bot)" | +fun bfilter :: "bexp \ bool \ 'a st option \ 'a st option" where +"bfilter (Bc v) res S = (if v=res then S else None)" | "bfilter (Not b) res S = bfilter b (\ res) S" | "bfilter (And b1 b2) res S = (if res then bfilter b1 True (bfilter b2 True S) else bfilter b1 False S \ bfilter b2 False S)" | "bfilter (Less e1 e2) res S = - (let (res1,res2) = filter_less' res (aval' e1 S) (aval' e2 S) + (let (res1,res2) = filter_less' res (aval'' e1 S) (aval'' e2 S) in afilter e1 res1 (afilter e2 res2 S))" -lemma afilter_sound: "s <:up S \ aval e s <: a \ s <:up afilter e a S" +lemma afilter_sound: "s : \\<^isub>u S \ aval e s : \ a \ s : \\<^isub>u (afilter e a S)" proof(induction e arbitrary: a S) case N thus ?case by simp next case (V x) - obtain S' where "S = Up S'" and "s <:f S'" using `s <:up S` + obtain S' where "S = Some S'" and "s : \\<^isub>f S'" using `s : \\<^isub>u S` by(auto simp: in_rep_up_iff) - moreover hence "s x <: lookup S' x" by(simp add: rep_st_def) - moreover have "s x <: a" using V by simp + moreover hence "s x : \ (lookup S' x)" by(simp add: rep_st_def) + moreover have "s x : \ a" using V by simp ultimately show ?case using V(1) by(simp add: lookup_update Let_def rep_st_def) - (metis le_rep emptyE in_rep_meet rep_Bot subset_empty) + (metis mono_rep emptyE in_rep_meet rep_Bot subset_empty) next case (Plus e1 e2) thus ?case - using filter_plus'[OF _ aval'_sound[OF Plus(3)] aval'_sound[OF Plus(3)]] + using filter_plus'[OF _ aval''_sound[OF Plus(3)] aval''_sound[OF Plus(3)]] by (auto split: prod.split) qed -lemma bfilter_sound: "s <:up S \ bv = bval b s \ s <:up bfilter b bv S" +lemma bfilter_sound: "s : \\<^isub>u S \ bv = bval b s \ s : \\<^isub>u(bfilter b bv S)" proof(induction b arbitrary: S bv) case Bc thus ?case by simp next @@ -138,16 +135,15 @@ next case (Less e1 e2) thus ?case by (auto split: prod.split) - (metis afilter_sound filter_less' aval'_sound Less) + (metis afilter_sound filter_less' aval''_sound Less) qed -fun step :: "'a st up \ 'a st up acom \ 'a st up acom" +fun step :: "'a st option \ 'a st option acom \ 'a st option acom" where "step S (SKIP {P}) = (SKIP {S})" | "step S (x ::= e {P}) = - x ::= e {case S of Bot \ Bot - | Up S \ Up(update S x (aval' e (Up S)))}" | + x ::= e {case S of None \ None | Some S \ Some(update S x (aval' e S))}" | "step S (c1; c2) = step S c1; step (post c1) c2" | "step S (IF b THEN c1 ELSE c2 {P}) = (let c1' = step (bfilter b True S) c1; c2' = step (bfilter b False S) c2 @@ -157,7 +153,7 @@ WHILE b DO step (bfilter b True Inv) c {bfilter b False Inv}" -definition AI :: "com \ 'a st up acom option" where +definition AI :: "com \ 'a st option acom option" where "AI = lpfp\<^isub>c (step \)" lemma strip_step[simp]: "strip(step S c) = strip c" @@ -166,68 +162,73 @@ subsubsection "Soundness" -lemma in_rep_update: "\ s <:f S; i <: a \ \ s(x := i) <:f update S x a" +lemma in_rep_update: + "\ s : \\<^isub>f S; i : \ a \ \ s(x := i) : \\<^isub>f(update S x a)" by(simp add: rep_st_def lookup_update) -lemma While_final_False: "(WHILE b DO c, s) \ t \ \ bval b t" -by(induct "WHILE b DO c" s t rule: big_step_induct) simp_all -lemma step_sound: - "step S c \ c \ (strip c,s) \ t \ s <:up S \ t <:up post c" -proof(induction c arbitrary: S s t) +lemma step_preserves_le2: + "\ S \ \\<^isub>u sa; cs \ \\<^isub>c ca; strip cs = c; strip ca = c \ + \ step_cs S cs \ \\<^isub>c (step sa ca)" +proof(induction c arbitrary: cs ca S sa) case SKIP thus ?case - by simp (metis skipE up_fun_in_rep_le) + by(auto simp:strip_eq_SKIP) next case Assign thus ?case - apply (auto simp del: fun_upd_apply split: up.splits) - by (metis aval'_sound fun_in_rep_le in_rep_update rep_up.simps(2)) + by (fastforce simp: strip_eq_Assign intro: aval'_sound in_rep_update + split: option.splits del:subsetD) next - case Semi thus ?case by simp blast + case Semi thus ?case apply (auto simp: strip_eq_Semi) + by (metis le_post post_map_acom) next - case (If b c1 c2 S0) - show ?case - proof cases - assume "bval b s" - with If.prems have 1: "step (bfilter b True S) c1 \ c1" - and 2: "(strip c1, s) \ t" and 3: "post c1 \ S0" by(auto simp: Let_def) - from If.IH(1)[OF 1 2 bfilter_sound[OF `s <:up S`]] `bval b s` 3 - show ?thesis by simp (metis up_fun_in_rep_le) - next - assume "\ bval b s" - with If.prems have 1: "step (bfilter b False S) c2 \ c2" - and 2: "(strip c2, s) \ t" and 3: "post c2 \ S0" by(auto simp: Let_def) - from If.IH(2)[OF 1 2 bfilter_sound[OF `s <:up S`]] `\ bval b s` 3 - show ?thesis by simp (metis up_fun_in_rep_le) - qed + case (If b c1 c2) + then obtain cs1 cs2 ca1 ca2 P Pa where + "cs= IF b THEN cs1 ELSE cs2 {P}" "ca= IF b THEN ca1 ELSE ca2 {Pa}" + "P \ \\<^isub>u Pa" "cs1 \ \\<^isub>c ca1" "cs2 \ \\<^isub>c ca2" + "strip cs1 = c1" "strip ca1 = c1" "strip cs2 = c2" "strip ca2 = c2" + by (fastforce simp: strip_eq_If) + moreover have "post cs1 \ \\<^isub>u(post ca1 \ post ca2)" + by (metis (no_types) `cs1 \ \\<^isub>c ca1` join_ge1 le_post mono_rep_u order_trans post_map_acom) + moreover have "post cs2 \ \\<^isub>u(post ca1 \ post ca2)" + by (metis (no_types) `cs2 \ \\<^isub>c ca2` join_ge2 le_post mono_rep_u order_trans post_map_acom) + ultimately show ?case using `S \ \\<^isub>u sa` + by (simp add: If.IH subset_iff bfilter_sound) next - case (While Inv b c P) - from While.prems have inv: "step (bfilter b True Inv) c \ c" - and "post c \ Inv" and "S \ Inv" and "bfilter b False Inv \ P" - by(auto simp: Let_def) - { fix s t have "(WHILE b DO strip c,s) \ t \ s <:up Inv \ t <:up Inv" - proof(induction "WHILE b DO strip c" s t rule: big_step_induct) - case WhileFalse thus ?case by simp - next - case (WhileTrue s1 s2 s3) - from WhileTrue.hyps(5)[OF up_fun_in_rep_le[OF While.IH[OF inv `(strip c, s1) \ s2` bfilter_sound[OF `s1 <:up Inv`]] `post c \ Inv`]] `bval b s1` - show ?case by simp - qed - } note Inv = this - from While.prems(2) have "(WHILE b DO strip c, s) \ t" and "\ bval b t" - by(auto dest: While_final_False) - from Inv[OF this(1) up_fun_in_rep_le[OF `s <:up S` `S \ Inv`]] - have "t <:up Inv" . - from up_fun_in_rep_le[OF bfilter_sound[OF this] `bfilter b False Inv \ P`] - show ?case using `\ bval b t` by simp + case (While b c1) + then obtain cs1 ca1 I P Ia Pa where + "cs = {I} WHILE b DO cs1 {P}" "ca = {Ia} WHILE b DO ca1 {Pa}" + "I \ \\<^isub>u Ia" "P \ \\<^isub>u Pa" "cs1 \ \\<^isub>c ca1" + "strip cs1 = c1" "strip ca1 = c1" + by (fastforce simp: strip_eq_While) + moreover have "S \ post cs1 \ \\<^isub>u (sa \ post ca1)" + using `S \ \\<^isub>u sa` le_post[OF `cs1 \ \\<^isub>c ca1`, simplified] + by (metis (no_types) join_ge1 join_ge2 le_sup_iff mono_rep_u order_trans) + ultimately show ?case by (simp add: While.IH subset_iff bfilter_sound) qed -lemma AI_sound: "\ AI c = Some c'; (c,s) \ t \ \ t <:up post c'" -unfolding AI_def -by (metis in_rep_Top_up lpfpc_pfp step_sound strip_lpfpc strip_step) -(* -by(metis step_sound[of "\" c' s t] strip_lpfp strip_step - lpfp_pfp mono_def mono_step[OF le_refl] in_rep_Top_up) -*) +lemma step_preserves_le: + "\ S \ \\<^isub>u sa; cs \ \\<^isub>c ca; strip cs = c \ + \ step_cs S cs \ \\<^isub>c(step sa ca)" +by (metis le_strip step_preserves_le2 strip_acom) + +lemma AI_sound: "AI c = Some c' \ CS UNIV c \ \\<^isub>c c'" +proof(simp add: CS_def AI_def) + assume 1: "lpfp\<^isub>c (step \) c = Some c'" + have 2: "step \ c' \ c'" by(rule lpfpc_pfp[OF 1]) + have 3: "strip (\\<^isub>c (step \ c')) = c" + by(simp add: strip_lpfpc[OF _ 1]) + have "lfp c (step_cs UNIV) \ \\<^isub>c (step \ c')" + proof(rule lfp_lowerbound[OF 3]) + show "step_cs UNIV (\\<^isub>c (step \ c')) \ \\<^isub>c (step \ c')" + proof(rule step_preserves_le[OF _ _ 3]) + show "UNIV \ \\<^isub>u \" by simp + show "\\<^isub>c (step \ c') \ \\<^isub>c c'" by(rule mono_rep_c[OF 2]) + qed + qed + from this 2 show "lfp c (step_cs UNIV) \ \\<^isub>c c'" + by (blast intro: mono_rep_c order_trans) +qed + end @@ -242,39 +243,44 @@ begin lemma mono_aval': "S \ S' \ aval' e S \ aval' e S'" +by(induction e) (auto simp: le_st_def lookup_def mono_plus') + +lemma mono_aval'': "S \ S' \ aval'' e S \ aval'' e S'" apply(cases S) apply simp apply(cases S') apply simp -apply simp -by(induction e) (auto simp: le_st_def lookup_def mono_plus') +by (simp add: mono_aval') lemma mono_afilter: "r \ r' \ S \ S' \ afilter e r S \ afilter e r' S'" apply(induction e arbitrary: r r' S S') -apply(auto simp: Let_def split: up.splits prod.splits) -apply (metis le_rep subsetD) +apply(auto simp: Let_def split: option.splits prod.splits) +apply (metis mono_rep subsetD) apply(drule_tac x = "list" in mono_lookup) apply (metis mono_meet le_trans) -apply (metis mono_meet mono_lookup mono_update le_trans) -apply(metis mono_aval' mono_filter_plus'[simplified le_prod_def] fst_conv snd_conv) +apply (metis mono_meet mono_lookup mono_update) +apply(metis mono_aval'' mono_filter_plus'[simplified le_prod_def] fst_conv snd_conv) done lemma mono_bfilter: "S \ S' \ bfilter b r S \ bfilter b r S'" apply(induction b arbitrary: r S S') apply(auto simp: le_trans[OF _ join_ge1] le_trans[OF _ join_ge2] split: prod.splits) -apply(metis mono_aval' mono_afilter mono_filter_less'[simplified le_prod_def] fst_conv snd_conv) +apply(metis mono_aval'' mono_afilter mono_filter_less'[simplified le_prod_def] fst_conv snd_conv) done lemma post_le_post: "c \ c' \ post c \ post c'" by (induction c c' rule: le_acom.induct) simp_all -lemma mono_step: "S \ S' \ c \ c' \ step S c \ step S' c'" +lemma mono_step_aux: "S \ S' \ c \ c' \ step S c \ step S' c'" apply(induction c c' arbitrary: S S' rule: le_acom.induct) apply (auto simp: post_le_post Let_def mono_bfilter mono_update mono_aval' le_join_disj - split: up.split) + split: option.split) done +lemma mono_step: "mono (step S)" +by(simp add: mono_def mono_step_aux[OF le_refl]) + end end