Theory Abs_Int1
section "Computable Abstract Interpretation"
theory Abs_Int1
imports Abs_State
begin
text‹Abstract interpretation over type ‹st› instead of
functions.›
context Gamma
begin
fun aval' :: "aexp ⇒ 'av st ⇒ 'av" where
"aval' (N n) S = num' n" |
"aval' (V x) S = lookup S x" |
"aval' (Plus a1 a2) S = plus' (aval' a1 S) (aval' a2 S)"
lemma aval'_sound: "s : γ⇩f S ⟹ aval a s : γ(aval' a S)"
by (induction a) (auto simp: gamma_num' gamma_plus' γ_st_def lookup_def)
end
text‹The for-clause (here and elsewhere) only serves the purpose of fixing
the name of the type parameter @{typ 'av} which would otherwise be renamed to
@{typ 'a}.›
locale Abs_Int = Gamma where γ=γ for γ :: "'av::SL_top ⇒ val set"
begin
fun step' :: "'av st option ⇒ 'av st option acom ⇒ 'av st option acom" where
"step' S (SKIP {P}) = (SKIP {S})" |
"step' S (x ::= e {P}) =
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' S c1; c2' = step' S c2
in IF b THEN c1' ELSE c2' {post c1 ⊔ post c2})" |
"step' S ({Inv} WHILE b DO c {P}) =
{S ⊔ post c} WHILE b DO step' Inv c {Inv}"
definition AI :: "com ⇒ 'av st option acom option" where
"AI = lpfp⇩c (step' ⊤)"
lemma strip_step'[simp]: "strip(step' S c) = strip c"
by(induct c arbitrary: S) (simp_all add: Let_def)
text‹Soundness:›
lemma in_gamma_update:
"⟦ s : γ⇩f S; i : γ a ⟧ ⟹ s(x := i) : γ⇩f(update S x a)"
by(simp add: γ_st_def lookup_update)
text‹The soundness proofs are textually identical to the ones for the step
function operating on states as functions.›
lemma step_preserves_le:
"⟦ S ⊆ γ⇩o S'; c ≤ γ⇩c c' ⟧ ⟹ step S c ≤ γ⇩c (step' S' c')"
proof(induction c arbitrary: c' S S')
case SKIP thus ?case by(auto simp:SKIP_le map_acom_SKIP)
next
case Assign thus ?case
by (fastforce simp: Assign_le map_acom_Assign intro: aval'_sound in_gamma_update
split: option.splits del:subsetD)
next
case Seq thus ?case apply (auto simp: Seq_le map_acom_Seq)
by (metis le_post post_map_acom)
next
case (If b c1 c2 P)
then obtain c1' c2' P' where
"c' = IF b THEN c1' ELSE c2' {P'}"
"P ⊆ γ⇩o P'" "c1 ≤ γ⇩c c1'" "c2 ≤ γ⇩c c2'"
by (fastforce simp: If_le map_acom_If)
moreover have "post c1 ⊆ γ⇩o(post c1' ⊔ post c2')"
by (metis (no_types) ‹c1 ≤ γ⇩c c1'› join_ge1 le_post mono_gamma_o order_trans post_map_acom)
moreover have "post c2 ⊆ γ⇩o(post c1' ⊔ post c2')"
by (metis (no_types) ‹c2 ≤ γ⇩c c2'› join_ge2 le_post mono_gamma_o order_trans post_map_acom)
ultimately show ?case using ‹S ⊆ γ⇩o S'› by (simp add: If.IH subset_iff)
next
case (While I b c1 P)
then obtain c1' I' P' where
"c' = {I'} WHILE b DO c1' {P'}"
"I ⊆ γ⇩o I'" "P ⊆ γ⇩o P'" "c1 ≤ γ⇩c c1'"
by (fastforce simp: map_acom_While While_le)
moreover have "S ∪ post c1 ⊆ γ⇩o (S' ⊔ post c1')"
using ‹S ⊆ γ⇩o S'› le_post[OF ‹c1 ≤ γ⇩c c1'›, simplified]
by (metis (no_types) join_ge1 join_ge2 le_sup_iff mono_gamma_o order_trans)
ultimately show ?case by (simp add: While.IH subset_iff)
qed
lemma AI_sound: "AI c = Some c' ⟹ CS c ≤ γ⇩c c'"
proof(simp add: CS_def AI_def)
assume 1: "lpfp⇩c (step' ⊤) c = Some c'"
have 2: "step' ⊤ c' ⊑ c'" by(rule lpfpc_pfp[OF 1])
have 3: "strip (γ⇩c (step' ⊤ c')) = c"
by(simp add: strip_lpfpc[OF _ 1])
have "lfp (step UNIV) c ≤ γ⇩c (step' ⊤ c')"
proof(rule lfp_lowerbound[simplified,OF 3])
show "step UNIV (γ⇩c (step' ⊤ c')) ≤ γ⇩c (step' ⊤ c')"
proof(rule step_preserves_le[OF _ _])
show "UNIV ⊆ γ⇩o ⊤" by simp
show "γ⇩c (step' ⊤ c') ≤ γ⇩c c'" by(rule mono_gamma_c[OF 2])
qed
qed
from this 2 show "lfp (step UNIV) c ≤ γ⇩c c'"
by (blast intro: mono_gamma_c order_trans)
qed
end
subsection "Monotonicity"
locale Abs_Int_mono = Abs_Int +
assumes mono_plus': "a1 ⊑ b1 ⟹ a2 ⊑ b2 ⟹ plus' a1 a2 ⊑ plus' b1 b2"
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_update: "a ⊑ a' ⟹ S ⊑ S' ⟹ update S x a ⊑ update S' x a'"
by(auto simp add: le_st_def lookup_def update_def)
lemma mono_step': "S ⊑ S' ⟹ c ⊑ c' ⟹ step' S c ⊑ step' S' c'"
apply(induction c c' arbitrary: S S' rule: le_acom.induct)
apply (auto simp: Let_def mono_update mono_aval' mono_post le_join_disj
split: option.split)
done
end
subsection "Ascending Chain Condition"
abbreviation "strict r == r ∩ -(r^-1)"
abbreviation "acc r == wf((strict r)^-1)"
lemma strict_inv_image: "strict(inv_image r f) = inv_image (strict r) f"
by(auto simp: inv_image_def)
lemma acc_inv_image:
"acc r ⟹ acc (inv_image r f)"
by (metis converse_inv_image strict_inv_image wf_inv_image)
text‹ACC for option type:›
lemma acc_option: assumes "acc {(x,y::'a::preord). x ⊑ y}"
shows "acc {(x,y::'a::preord option). x ⊑ y}"
proof(auto simp: wf_eq_minimal)
fix xo :: "'a option" and Qo assume "xo : Qo"
let ?Q = "{x. Some x ∈ Qo}"
show "∃yo∈Qo. ∀zo. yo ⊑ zo ∧ ~ zo ⊑ yo ⟶ zo ∉ Qo" (is "∃zo∈Qo. ?P zo")
proof cases
assume "?Q = {}"
hence "?P None" by auto
moreover have "None ∈ Qo" using ‹?Q = {}› ‹xo : Qo›
by auto (metis not_Some_eq)
ultimately show ?thesis by blast
next
assume "?Q ≠ {}"
with assms show ?thesis
apply(auto simp: wf_eq_minimal)
apply(erule_tac x="?Q" in allE)
apply auto
apply(rule_tac x = "Some z" in bexI)
by auto
qed
qed
text‹ACC for abstract states, via measure functions.›
lemma measure_st: assumes "(strict{(x,y::'a::SL_top). x ⊑ y})^-1 <= measure m"
and "∀x y::'a::SL_top. x ⊑ y ∧ y ⊑ x ⟶ m x = m y"
shows "(strict{(S,S'::'a::SL_top st). S ⊑ S'})^-1 ⊆
measure(%fd. ∑x| x∈set(dom fd) ∧ ~ ⊤ ⊑ fun fd x. m(fun fd x)+1)"
proof-
{ fix S S' :: "'a st" assume "S ⊑ S'" "~ S' ⊑ S"
let ?X = "set(dom S)" let ?Y = "set(dom S')"
let ?f = "fun S" let ?g = "fun S'"
let ?X' = "{x:?X. ~ ⊤ ⊑ ?f x}" let ?Y' = "{y:?Y. ~ ⊤ ⊑ ?g y}"
from ‹S ⊑ S'› have "∀y∈?Y'∩?X. ?f y ⊑ ?g y"
by(auto simp: le_st_def lookup_def)
hence 1: "∀y∈?Y'∩?X. m(?g y)+1 ≤ m(?f y)+1"
using assms(1,2) by(fastforce)
from ‹~ S' ⊑ S› obtain u where u: "u : ?X" "~ lookup S' u ⊑ ?f u"
by(auto simp: le_st_def)
hence "u : ?X'" by simp (metis preord_class.le_trans top)
have "?Y'-?X = {}" using ‹S ⊑ S'› by(fastforce simp: le_st_def lookup_def)
have "?Y'∩?X <= ?X'" apply auto
apply (metis ‹S ⊑ S'› le_st_def lookup_def preord_class.le_trans)
done
have "(∑y∈?Y'. m(?g y)+1) = (∑y∈(?Y'-?X) ∪ (?Y'∩?X). m(?g y)+1)"
by (metis Un_Diff_Int)
also have "… = (∑y∈?Y'∩?X. m(?g y)+1)"
using ‹?Y'-?X = {}› by (metis Un_empty_left)
also have "… < (∑x∈?X'. m(?f x)+1)"
proof cases
assume "u ∈ ?Y'"
hence "m(?g u) < m(?f u)" using assms(1) ‹S ⊑ S'› u
by (fastforce simp: le_st_def lookup_def)
have "(∑y∈?Y'∩?X. m(?g y)+1) < (∑y∈?Y'∩?X. m(?f y)+1)"
using ‹u:?X› ‹u:?Y'› ‹m(?g u) < m(?f u)›
by(fastforce intro!: sum_strict_mono_ex1[OF _ 1])
also have "… ≤ (∑y∈?X'. m(?f y)+1)"
by(simp add: sum_mono2[OF _ ‹?Y'∩?X <= ?X'›])
finally show ?thesis .
next
assume "u ∉ ?Y'"
with ‹?Y'∩?X <= ?X'› have "?Y'∩?X - {u} <= ?X' - {u}" by blast
have "(∑y∈?Y'∩?X. m(?g y)+1) = (∑y∈?Y'∩?X - {u}. m(?g y)+1)"
proof-
have "?Y'∩?X = ?Y'∩?X - {u}" using ‹u ∉ ?Y'› by auto
thus ?thesis by metis
qed
also have "… < (∑y∈?Y'∩?X-{u}. m(?g y)+1) + (∑y∈{u}. m(?f y)+1)" by simp
also have "(∑y∈?Y'∩?X-{u}. m(?g y)+1) ≤ (∑y∈?Y'∩?X-{u}. m(?f y)+1)"
using 1 by(blast intro: sum_mono)
also have "… ≤ (∑y∈?X'-{u}. m(?f y)+1)"
by(simp add: sum_mono2[OF _ ‹?Y'∩?X-{u} <= ?X'-{u}›])
also have "… + (∑y∈{u}. m(?f y)+1)= (∑y∈(?X'-{u}) ∪ {u}. m(?f y)+1)"
using ‹u:?X'› by(subst sum.union_disjoint[symmetric]) auto
also have "… = (∑x∈?X'. m(?f x)+1)"
using ‹u : ?X'› by(simp add:insert_absorb)
finally show ?thesis by (blast intro: add_right_mono)
qed
finally have "(∑y∈?Y'. m(?g y)+1) < (∑x∈?X'. m(?f x)+1)" .
} thus ?thesis by(auto simp add: measure_def inv_image_def)
qed
text‹ACC for acom. First the ordering on acom is related to an ordering on
lists of annotations.›
lemma listrel_Cons_iff:
"(x#xs, y#ys) : listrel r ⟷ (x,y) ∈ r ∧ (xs,ys) ∈ listrel r"
by (blast intro:listrel.Cons)
lemma listrel_app: "(xs1,ys1) : listrel r ⟹ (xs2,ys2) : listrel r
⟹ (xs1@xs2, ys1@ys2) : listrel r"
by(auto simp add: listrel_iff_zip)
lemma listrel_app_same_size: "size xs1 = size ys1 ⟹ size xs2 = size ys2 ⟹
(xs1@xs2, ys1@ys2) : listrel r ⟷
(xs1,ys1) : listrel r ∧ (xs2,ys2) : listrel r"
by(auto simp add: listrel_iff_zip)
lemma listrel_converse: "listrel(r^-1) = (listrel r)^-1"
proof-
{ fix xs ys
have "(xs,ys) : listrel(r^-1) ⟷ (ys,xs) : listrel r"
apply(induct xs arbitrary: ys)
apply (fastforce simp: listrel.Nil)
apply (fastforce simp: listrel_Cons_iff)
done
} thus ?thesis by auto
qed
lemma acc_listrel: fixes r :: "('a*'a)set" assumes "refl r" and "trans r"
and "acc r" shows "acc (listrel r - {([],[])})"
proof-
have refl: "!!x. (x,x) : r" using ‹refl r› unfolding refl_on_def by blast
have trans: "!!x y z. (x,y) : r ⟹ (y,z) : r ⟹ (x,z) : r"
using ‹trans r› unfolding trans_def by blast
from assms(3) obtain mx :: "'a set ⇒ 'a" where
mx: "!!S x. x:S ⟹ mx S : S ∧ (∀y. (mx S,y) : strict r ⟶ y ∉ S)"
by(simp add: wf_eq_minimal) metis
let ?R = "listrel r - {([], [])}"
{ fix Q and xs :: "'a list"
have "xs ∈ Q ⟹ ∃ys. ys∈Q ∧ (∀zs. (ys, zs) ∈ strict ?R ⟶ zs ∉ Q)"
(is "_ ⟹ ∃ys. ?P Q ys")
proof(induction xs arbitrary: Q rule: length_induct)
case (1 xs)
{ have "!!ys Q. size ys < size xs ⟹ ys : Q ⟹ ∃ms. ?P Q ms"
using "1.IH" by blast
} note IH = this
show ?case
proof(cases xs)
case Nil with ‹xs : Q› have "?P Q []" by auto
thus ?thesis by blast
next
case (Cons x ys)
let ?Q1 = "{a. ∃bs. size bs = size ys ∧ a#bs : Q}"
have "x : ?Q1" using ‹xs : Q› Cons by auto
from mx[OF this] obtain m1 where
1: "m1 ∈ ?Q1 ∧ (∀y. (m1,y) ∈ strict r ⟶ y ∉ ?Q1)" by blast
then obtain ms1 where "size ms1 = size ys" "m1#ms1 : Q" by blast+
hence "size ms1 < size xs" using Cons by auto
let ?Q2 = "{bs. ∃m1'. (m1',m1):r ∧ (m1,m1'):r ∧ m1'#bs : Q ∧ size bs = size ms1}"
have "ms1 : ?Q2" using ‹m1#ms1 : Q› by(blast intro: refl)
from IH[OF ‹size ms1 < size xs› this]
obtain ms where 2: "?P ?Q2 ms" by auto
then obtain m1' where m1': "(m1',m1) : r ∧ (m1,m1') : r ∧ m1'#ms : Q"
by blast
hence "∀ab. (m1'#ms,ab) : strict ?R ⟶ ab ∉ Q" using 1 2
apply (auto simp: listrel_Cons_iff)
apply (metis ‹length ms1 = length ys› listrel_eq_len trans)
by (metis ‹length ms1 = length ys› listrel_eq_len trans)
with m1' show ?thesis by blast
qed
qed
}
thus ?thesis unfolding wf_eq_minimal by (metis converse_iff)
qed
lemma le_iff_le_annos: "c1 ⊑ c2 ⟷
(annos c1, annos c2) : listrel{(x,y). x ⊑ y} ∧ strip c1 = strip c2"
apply(induct c1 c2 rule: le_acom.induct)
apply (auto simp: listrel.Nil listrel_Cons_iff listrel_app size_annos_same2)
apply (metis listrel_app_same_size size_annos_same)+
done
lemma le_acom_subset_same_annos:
"(strict{(c,c'::'a::preord acom). c ⊑ c'})^-1 ⊆
(strict(inv_image (listrel{(a,a'::'a). a ⊑ a'} - {([],[])}) annos))^-1"
by(auto simp: le_iff_le_annos)
lemma acc_acom: "acc {(a,a'::'a::preord). a ⊑ a'} ⟹
acc {(c,c'::'a acom). c ⊑ c'}"
apply(rule wf_subset[OF _ le_acom_subset_same_annos])
apply(rule acc_inv_image[OF acc_listrel])
apply(auto simp: refl_on_def trans_def intro: le_trans)
done
text‹Termination of the fixed-point finders, assuming monotone functions:›
lemma pfp_termination:
fixes x0 :: "'a::preord"
assumes mono: "⋀x y. x ⊑ y ⟹ f x ⊑ f y" and "acc {(x::'a,y). x ⊑ y}"
and "x0 ⊑ f x0" shows "∃x. pfp f x0 = Some x"
proof(simp add: pfp_def, rule wf_while_option_Some[where P = "%x. x ⊑ f x"])
show "wf {(x, s). (s ⊑ f s ∧ ¬ f s ⊑ s) ∧ x = f s}"
by(rule wf_subset[OF assms(2)]) auto
next
show "x0 ⊑ f x0" by(rule assms)
next
fix x assume "x ⊑ f x" thus "f x ⊑ f(f x)" by(rule mono)
qed
lemma lpfpc_termination:
fixes f :: "(('a::SL_top)option acom ⇒ 'a option acom)"
assumes "acc {(x::'a,y). x ⊑ y}" and "⋀x y. x ⊑ y ⟹ f x ⊑ f y"
and "⋀c. strip(f c) = strip c"
shows "∃c'. lpfp⇩c f c = Some c'"
unfolding lpfp⇩c_def
apply(rule pfp_termination)
apply(erule assms(2))
apply(rule acc_acom[OF acc_option[OF assms(1)]])
apply(simp add: bot_acom assms(3))
done
context Abs_Int_mono
begin
lemma AI_Some_measure:
assumes "(strict{(x,y::'a). x ⊑ y})^-1 <= measure m"
and "∀x y::'a. x ⊑ y ∧ y ⊑ x ⟶ m x = m y"
shows "∃c'. AI c = Some c'"
unfolding AI_def
apply(rule lpfpc_termination)
apply(rule wf_subset[OF wf_measure measure_st[OF assms]])
apply(erule mono_step'[OF le_refl])
apply(rule strip_step')
done
end
end