Theory Abs_Int0

(* Author: Tobias Nipkow *)

section "Abstract Interpretation Abstractly"

theory Abs_Int0
imports
  "HOL-Library.While_Combinator"
  Collecting
begin

subsection "Orderings"

class preord =
fixes le :: "'a  'a  bool" (infix "" 50)
assumes le_refl[simp]: "x  x"
and le_trans: "x  y  y  z  x  z"
begin

definition mono where "mono f = (x y. x  y  f x  f y)"

lemma monoD: "mono f  x  y  f x  f y" by(simp add: mono_def)

lemma mono_comp: "mono f  mono g  mono (g o f)"
by(simp add: mono_def)

declare le_trans[trans]

end

text‹Note: no antisymmetry. Allows implementations where some abstract
element is implemented by two different values @{prop "x  y"}
such that @{prop"x  y"} and @{prop"y  x"}. Antisymmetry is not
needed because we never compare elements for equality but only for ⊑›.
›

class SL_top = preord +
fixes join :: "'a  'a  'a" (infixl "" 65)
fixes Top :: "'a" ("")
assumes join_ge1 [simp]: "x  x  y"
and join_ge2 [simp]: "y  x  y"
and join_least: "x  z  y  z  x  y  z"
and top[simp]: "x  "
begin

lemma join_le_iff[simp]: "x  y  z  x  z  y  z"
by (metis join_ge1 join_ge2 join_least le_trans)

lemma le_join_disj: "x  y  x  z  x  y  z"
by (metis join_ge1 join_ge2 le_trans)

end

instantiation "fun" :: (type, SL_top) SL_top
begin

definition "f  g = (x. f x  g x)"
definition "f  g = (λx. f x  g x)"
definition " = (λx. )"

lemma join_apply[simp]: "(f  g) x = f x  g x"
by (simp add: join_fun_def)

instance
proof (standard,goal_cases)
  case 2 thus ?case by (metis le_fun_def preord_class.le_trans)
qed (simp_all add: le_fun_def Top_fun_def)

end


instantiation acom :: (preord) preord
begin

fun le_acom :: "('a::preord)acom  'a acom  bool" where
"le_acom (SKIP {S}) (SKIP {S'}) = (S  S')" |
"le_acom (x ::= e {S}) (x' ::= e' {S'}) = (x=x'  e=e'  S  S')" |
"le_acom (c1;;c2) (c1';;c2') = (le_acom c1 c1'  le_acom c2 c2')" |
"le_acom (IF b THEN c1 ELSE c2 {S}) (IF b' THEN c1' ELSE c2' {S'}) =
  (b=b'  le_acom c1 c1'  le_acom c2 c2'  S  S')" |
"le_acom ({Inv} WHILE b DO c {P}) ({Inv'} WHILE b' DO c' {P'}) =
  (b=b'  le_acom c c'  Inv  Inv'  P  P')" |
"le_acom _ _ = False"

lemma [simp]: "SKIP {S}  c  (S'. c = SKIP {S'}  S  S')"
by (cases c) auto

lemma [simp]: "x ::= e {S}  c  (S'. c = x ::= e {S'}  S  S')"
by (cases c) auto

lemma [simp]: "c1;;c2  c  (c1' c2'. c = c1';;c2'  c1  c1'  c2  c2')"
by (cases c) auto

lemma [simp]: "IF b THEN c1 ELSE c2 {S}  c 
  (c1' c2' S'. c = IF b THEN c1' ELSE c2' {S'}  c1  c1'  c2  c2'  S  S')"
by (cases c) auto

lemma [simp]: "{Inv} WHILE b DO c {P}  w 
  (Inv' c' P'. w = {Inv'} WHILE b DO c' {P'}  c  c'  Inv  Inv'  P  P')"
by (cases w) auto

instance
proof (standard,goal_cases)
  case (1 x) thus ?case by (induct x) auto
next
  case (2 x y z) thus ?case
  apply(induct x y arbitrary: z rule: le_acom.induct)
  apply (auto intro: le_trans)
  done
qed

end


subsubsection "Lifting"

instantiation option :: (preord)preord
begin

fun le_option where
"Some x  Some y = (x  y)" |
"None  y = True" |
"Some _  None = False"

lemma [simp]: "(x  None) = (x = None)"
by (cases x) simp_all

lemma [simp]: "(Some x  u) = (y. u = Some y & x  y)"
by (cases u) auto

instance
proof (standard,goal_cases)
  case (1 x) show ?case by(cases x, simp_all)
next
  case (2 x y z) thus ?case
    by(cases z, simp, cases y, simp, cases x, auto intro: le_trans)
qed

end

instantiation option :: (SL_top)SL_top
begin

fun join_option where
"Some x  Some y = Some(x  y)" |
"None  y = y" |
"x  None = x"

lemma join_None2[simp]: "x  None = x"
by (cases x) simp_all

definition " = Some "

instance
proof (standard,goal_cases)
  case (1 x y) thus ?case by(cases x, simp, cases y, simp_all)
next
  case (2 x y) thus ?case by(cases y, simp, cases x, simp_all)
next
  case (3 x y z) thus ?case by(cases z, simp, cases y, simp, cases x, simp_all)
next
  case (4 x) thus ?case by(cases x, simp_all add: Top_option_def)
qed

end

definition bot_acom :: "com  ('a::SL_top)option acom" ("c") where
"c = anno None"

lemma strip_bot_acom[simp]: "strip(c c) = c"
by(simp add: bot_acom_def)

lemma bot_acom[rule_format]: "strip c' = c  c c  c'"
apply(induct c arbitrary: c')
apply (simp_all add: bot_acom_def)
 apply(induct_tac c')
  apply simp_all
 apply(induct_tac c')
  apply simp_all
 apply(induct_tac c')
  apply simp_all
 apply(induct_tac c')
  apply simp_all
 apply(induct_tac c')
  apply simp_all
done


subsubsection "Post-fixed point iteration"

definition
  pfp :: "(('a::preord)  'a)  'a  'a option" where
"pfp f = while_option (λx. ¬ f x  x) f"

lemma pfp_pfp: assumes "pfp f x0 = Some x" shows "f x  x"
using while_option_stop[OF assms[simplified pfp_def]] by simp

lemma pfp_least:
assumes mono: "x y. x  y  f x  f y"
and "f p  p" and "x0  p" and "pfp f x0 = Some x" shows "x  p"
proof-
  { fix x assume "x  p"
    hence  "f x  f p" by(rule mono)
    from this f p  p have "f x  p" by(rule le_trans)
  }
  thus "x  p" using assms(2-) while_option_rule[where P = "%x. x  p"]
    unfolding pfp_def by blast
qed

definition
 lpfpc :: "(('a::SL_top)option acom  'a option acom)  com  'a option acom option" where
"lpfpc f c = pfp f (c c)"

lemma lpfpc_pfp: "lpfpc f c0 = Some c  f c  c"
by(simp add: pfp_pfp lpfpc_def)

lemma strip_pfp:
assumes "x. g(f x) = g x" and "pfp f x0 = Some x" shows "g x = g x0"
using assms while_option_rule[where P = "%x. g x = g x0" and c = f]
unfolding pfp_def by metis

lemma strip_lpfpc: assumes "c. strip(f c) = strip c" and "lpfpc f c = Some c'"
shows "strip c' = c"
using assms(1) strip_pfp[OF _ assms(2)[simplified lpfpc_def]]
by(metis strip_bot_acom)

lemma lpfpc_least:
assumes mono: "x y. x  y  f x  f y"
and "strip p = c0" and "f p  p" and lp: "lpfpc f c0 = Some c" shows "c  p"
using pfp_least[OF _ _ bot_acom[OF strip p = c0] lp[simplified lpfpc_def]]
  mono f p  p
by blast


subsection "Abstract Interpretation"

definition γ_fun :: "('a  'b set)  ('c  'a)  ('c  'b)set" where
"γ_fun γ F = {f. x. f x  γ(F x)}"

fun γ_option :: "('a  'b set)  'a option  'b set" where
"γ_option γ None = {}" |
"γ_option γ (Some a) = γ a"

text‹The interface for abstract values:›

locale Val_abs =
fixes γ :: "'av::SL_top  val set"
  assumes mono_gamma: "a  b  γ a  γ b"
  and gamma_Top[simp]: "γ  = UNIV"
fixes num' :: "val  'av"
and plus' :: "'av  'av  'av"
  assumes gamma_num': "n : γ(num' n)"
  and gamma_plus':
 "n1 : γ a1  n2 : γ a2  n1+n2 : γ(plus' a1 a2)"

type_synonym 'av st = "(vname  'av)"

locale Abs_Int_Fun = Val_abs γ for γ :: "'av::SL_top  val set"
begin

fun aval' :: "aexp  'av st  'av" where
"aval' (N n) S = num' n" |
"aval' (V x) S = S x" |
"aval' (Plus a1 a2) S = plus' (aval' a1 S) (aval' a2 S)"

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(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}) =
   IF b THEN step' S c1 ELSE step' S 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 = lpfpc (step' )"


lemma strip_step'[simp]: "strip(step' S c) = strip c"
by(induct c arbitrary: S) (simp_all add: Let_def)


abbreviation γf :: "'av st  state set"
where "γf == γ_fun γ"

abbreviation γo :: "'av st option  state set"
where "γo == γ_option γf"

abbreviation γc :: "'av st option acom  state set acom"
where "γc == map_acom γo"

lemma gamma_f_Top[simp]: "γf Top = UNIV"
by(simp add: Top_fun_def γ_fun_def)

lemma gamma_o_Top[simp]: "γo Top = UNIV"
by (simp add: Top_option_def)

(* FIXME (maybe also le → sqle?) *)

lemma mono_gamma_f: "f  g  γf f  γf g"
by(auto simp: le_fun_def γ_fun_def dest: mono_gamma)

lemma mono_gamma_o:
  "sa  sa'  γo sa  γo sa'"
by(induction sa sa' rule: le_option.induct)(simp_all add: mono_gamma_f)

lemma mono_gamma_c: "ca  ca'  γc ca  γc ca'"
by (induction ca ca' rule: le_acom.induct) (simp_all add:mono_gamma_o)

text‹Soundness:›

lemma aval'_sound: "s : γf S  aval a s : γ(aval' a S)"
by (induct a) (auto simp: gamma_num' gamma_plus' γ_fun_def)

lemma in_gamma_update:
  " s : γf S; i : γ a   s(x := i) : γf(S(x := a))"
by(simp add: γ_fun_def)

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: "lpfpc (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
  with 2 show "lfp (step UNIV) c  γc c'"
    by (blast intro: mono_gamma_c order_trans)
qed

end


subsubsection "Monotonicity"

lemma mono_post: "c  c'  post c  post c'"
by(induction c c' rule: le_acom.induct) (auto)

locale Abs_Int_Fun_mono = Abs_Int_Fun +
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_fun_def mono_plus')

lemma mono_update: "a  a'  S  S'  S(x := a)  S'(x := a')"
by(simp add: le_fun_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

text‹Problem: not executable because of the comparison of abstract states,
i.e. functions, in the post-fixedpoint computation.›

end