src/HOL/IMP/Abs_Int_Den/Abs_Int_den1.thy
author nipkow
Fri, 17 May 2013 08:19:52 +0200
changeset 52046 bc01725d7918
parent 47818 151d137f1095
child 53015 a1119cf551e8
permissions -rw-r--r--
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
     1
(* Author: Tobias Nipkow *)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
     2
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents: 45110
diff changeset
     3
theory Abs_Int_den1
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents: 45110
diff changeset
     4
imports Abs_Int_den0_const
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
     5
begin
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
     6
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
     7
subsection "Backward Analysis of Expressions"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
     8
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
     9
class L_top_bot = SL_top +
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    10
fixes meet :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<sqinter>" 65)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    11
and Bot :: "'a"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    12
assumes meet_le1 [simp]: "x \<sqinter> y \<sqsubseteq> x"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    13
and meet_le2 [simp]: "x \<sqinter> y \<sqsubseteq> y"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    14
and meet_greatest: "x \<sqsubseteq> y \<Longrightarrow> x \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> y \<sqinter> z"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    15
assumes bot[simp]: "Bot \<sqsubseteq> x"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    16
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    17
locale Rep1 = Rep rep for rep :: "'a::L_top_bot \<Rightarrow> 'b set" +
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    18
assumes inter_rep_subset_rep_meet: "rep a1 \<inter> rep a2 \<subseteq> rep(a1 \<sqinter> a2)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    19
and rep_Bot: "rep Bot = {}"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    20
begin
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    21
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    22
lemma in_rep_meet: "x <: a1 \<Longrightarrow> x <: a2 \<Longrightarrow> x <: a1 \<sqinter> a2"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    23
by (metis IntI inter_rep_subset_rep_meet set_mp)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    24
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    25
lemma rep_meet[simp]: "rep(a1 \<sqinter> a2) = rep a1 \<inter> rep a2"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    26
by (metis equalityI inter_rep_subset_rep_meet le_inf_iff le_rep meet_le1 meet_le2)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    27
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    28
end
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    29
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    30
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    31
locale Val_abs1 = Val_abs rep num' plus' + Rep1 rep
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    32
  for rep :: "'a::L_top_bot \<Rightarrow> int set" and num' plus' +
45023
76abd26e2e2d renamed inv -> filter
nipkow
parents: 45017
diff changeset
    33
fixes filter_plus' :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a * 'a"
76abd26e2e2d renamed inv -> filter
nipkow
parents: 45017
diff changeset
    34
and filter_less' :: "bool \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a * 'a"
76abd26e2e2d renamed inv -> filter
nipkow
parents: 45017
diff changeset
    35
assumes filter_plus': "filter_plus' a a1 a2 = (a1',a2') \<Longrightarrow>
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    36
  n1 <: a1 \<Longrightarrow> n2 <: a2 \<Longrightarrow> n1+n2 <: a \<Longrightarrow> n1 <: a1' \<and> n2 <: a2'"
45023
76abd26e2e2d renamed inv -> filter
nipkow
parents: 45017
diff changeset
    37
and filter_less': "filter_less' (n1<n2) a1 a2 = (a1',a2') \<Longrightarrow>
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    38
  n1 <: a1 \<Longrightarrow> n2 <: a2 \<Longrightarrow> n1 <: a1' \<and> n2 <: a2'"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    39
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    40
datatype 'a up = bot | Up 'a
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    41
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    42
instantiation up :: (SL_top)SL_top
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    43
begin
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    44
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    45
fun le_up where
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    46
"Up x \<sqsubseteq> Up y = (x \<sqsubseteq> y)" |
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    47
"bot \<sqsubseteq> y = True" |
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    48
"Up _ \<sqsubseteq> bot = False"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    49
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    50
lemma [simp]: "(x \<sqsubseteq> bot) = (x = bot)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    51
by (cases x) simp_all
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    52
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    53
lemma [simp]: "(Up x \<sqsubseteq> u) = (EX y. u = Up y & x \<sqsubseteq> y)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    54
by (cases u) auto
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    55
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    56
fun join_up where
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    57
"Up x \<squnion> Up y = Up(x \<squnion> y)" |
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    58
"bot \<squnion> y = y" |
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    59
"x \<squnion> bot = x"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    60
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    61
lemma [simp]: "x \<squnion> bot = x"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    62
by (cases x) simp_all
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    63
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    64
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    65
definition "Top = Up Top"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    66
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    67
instance proof
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    68
  case goal1 show ?case by(cases x, simp_all)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    69
next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    70
  case goal2 thus ?case
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    71
    by(cases z, simp, cases y, simp, cases x, auto intro: le_trans)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    72
next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    73
  case goal3 thus ?case by(cases x, simp, cases y, simp_all)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    74
next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    75
  case goal4 thus ?case by(cases y, simp, cases x, simp_all)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    76
next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    77
  case goal5 thus ?case by(cases z, simp, cases y, simp, cases x, simp_all)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    78
next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    79
  case goal6 thus ?case by(cases x, simp_all add: Top_up_def)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    80
qed
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    81
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    82
end
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    83
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    84
44932
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents: 44923
diff changeset
    85
locale Abs_Int1 = Val_abs1 +
44944
f136409c2cef tuned post fixpoint setup
nipkow
parents: 44932
diff changeset
    86
fixes pfp :: "('a astate up \<Rightarrow> 'a astate up) \<Rightarrow> 'a astate up \<Rightarrow> 'a astate up"
f136409c2cef tuned post fixpoint setup
nipkow
parents: 44932
diff changeset
    87
assumes pfp: "f(pfp f x0) \<sqsubseteq> pfp f x0"
f136409c2cef tuned post fixpoint setup
nipkow
parents: 44932
diff changeset
    88
assumes above: "x0 \<sqsubseteq> pfp f x0"
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    89
begin
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    90
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    91
(* FIXME avoid duplicating this defn *)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    92
abbreviation astate_in_rep (infix "<:" 50) where
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    93
"s <: S == ALL x. s x <: lookup S x"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    94
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    95
abbreviation in_rep_up :: "state \<Rightarrow> 'a astate up \<Rightarrow> bool"  (infix "<::" 50) where
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    96
"s <:: S == EX S0. S = Up S0 \<and> s <: S0"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    97
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    98
lemma in_rep_up_trans: "(s::state) <:: S \<Longrightarrow> S \<sqsubseteq> T \<Longrightarrow> s <:: T"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    99
apply auto
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   100
by (metis in_mono le_astate_def le_rep lookup_def top)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   101
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   102
lemma in_rep_join_UpI: "s <:: S1 | s <:: S2 \<Longrightarrow> s <:: S1 \<squnion> S2"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   103
by (metis in_rep_up_trans SL_top_class.join_ge1 SL_top_class.join_ge2)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   104
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   105
fun aval' :: "aexp \<Rightarrow> 'a astate up \<Rightarrow> 'a" ("aval\<^isup>#") where
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   106
"aval' _ bot = Bot" |
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   107
"aval' (N n) _ = num' n" |
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   108
"aval' (V x) (Up S) = lookup S x" |
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   109
"aval' (Plus a1 a2) S = plus' (aval' a1 S) (aval' a2 S)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   110
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   111
lemma aval'_sound: "s <:: S \<Longrightarrow> aval a s <: aval' a S"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   112
by (induct a) (auto simp: rep_num' rep_plus')
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   113
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   114
fun afilter :: "aexp \<Rightarrow> 'a \<Rightarrow> 'a astate up \<Rightarrow> 'a astate up" where
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   115
"afilter (N n) a S = (if n <: a then S else bot)" |
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   116
"afilter (V x) a S = (case S of bot \<Rightarrow> bot | Up S \<Rightarrow>
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   117
  let a' = lookup S x \<sqinter> a in
44932
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents: 44923
diff changeset
   118
  if a' \<sqsubseteq> Bot then bot else Up(update S x a'))" |
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   119
"afilter (Plus e1 e2) a S =
45023
76abd26e2e2d renamed inv -> filter
nipkow
parents: 45017
diff changeset
   120
 (let (a1,a2) = filter_plus' a (aval' e1 S) (aval' e2 S)
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   121
  in afilter e1 a1 (afilter e2 a2 S))"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   122
44932
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents: 44923
diff changeset
   123
text{* The test for @{const Bot} in the @{const V}-case is important: @{const
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents: 44923
diff changeset
   124
Bot} indicates that a variable has no possible values, i.e.\ that the current
45017
07a0638c351a fixed two typos in IMP (by Jean Pichon)
kleing
parents: 45015
diff changeset
   125
program point is unreachable. But then the abstract state should collapse to
44932
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents: 44923
diff changeset
   126
@{const bot}. Put differently, we maintain the invariant that in an abstract
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents: 44923
diff changeset
   127
state all variables are mapped to non-@{const Bot} values. Otherwise the
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents: 44923
diff changeset
   128
(pointwise) join of two abstract states, one of which contains @{const Bot}
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents: 44923
diff changeset
   129
values, may produce too large a result, thus making the analysis less
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents: 44923
diff changeset
   130
precise. *}
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents: 44923
diff changeset
   131
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents: 44923
diff changeset
   132
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   133
fun bfilter :: "bexp \<Rightarrow> bool \<Rightarrow> 'a astate up \<Rightarrow> 'a astate up" where
45200
1f1897ac7877 renamed B to Bc
nipkow
parents: 45111
diff changeset
   134
"bfilter (Bc v) res S = (if v=res then S else bot)" |
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   135
"bfilter (Not b) res S = bfilter b (\<not> res) S" |
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   136
"bfilter (And b1 b2) res S =
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   137
  (if res then bfilter b1 True (bfilter b2 True S)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   138
   else bfilter b1 False S \<squnion> bfilter b2 False S)" |
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   139
"bfilter (Less e1 e2) res S =
45023
76abd26e2e2d renamed inv -> filter
nipkow
parents: 45017
diff changeset
   140
  (let (res1,res2) = filter_less' res (aval' e1 S) (aval' e2 S)
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   141
   in afilter e1 res1 (afilter e2 res2 S))"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   142
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   143
lemma afilter_sound: "s <:: S \<Longrightarrow> aval e s <: a \<Longrightarrow> s <:: afilter e a S"
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44944
diff changeset
   144
proof(induction e arbitrary: a S)
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   145
  case N thus ?case by simp
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   146
next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   147
  case (V x)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   148
  obtain S' where "S = Up S'" and "s <: S'" using `s <:: S` by auto
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   149
  moreover hence "s x <: lookup S' x" by(simp)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   150
  moreover have "s x <: a" using V by simp
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   151
  ultimately show ?case using V(1)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   152
    by(simp add: lookup_update Let_def)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   153
       (metis le_rep emptyE in_rep_meet rep_Bot subset_empty)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   154
next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   155
  case (Plus e1 e2) thus ?case
45023
76abd26e2e2d renamed inv -> filter
nipkow
parents: 45017
diff changeset
   156
    using filter_plus'[OF _ aval'_sound[OF Plus(3)] aval'_sound[OF Plus(3)]]
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   157
    by (auto split: prod.split)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   158
qed
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   159
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   160
lemma bfilter_sound: "s <:: S \<Longrightarrow> bv = bval b s \<Longrightarrow> s <:: bfilter b bv S"
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44944
diff changeset
   161
proof(induction b arbitrary: S bv)
45200
1f1897ac7877 renamed B to Bc
nipkow
parents: 45111
diff changeset
   162
  case Bc thus ?case by simp
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   163
next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   164
  case (Not b) thus ?case by simp
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   165
next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   166
  case (And b1 b2) thus ?case by (auto simp: in_rep_join_UpI)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   167
next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   168
  case (Less e1 e2) thus ?case
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   169
    by (auto split: prod.split)
45023
76abd26e2e2d renamed inv -> filter
nipkow
parents: 45017
diff changeset
   170
       (metis afilter_sound filter_less' aval'_sound Less)
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   171
qed
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   172
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   173
fun AI :: "com \<Rightarrow> 'a astate up \<Rightarrow> 'a astate up" where
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   174
"AI SKIP S = S" |
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   175
"AI (x ::= a) S =
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   176
  (case S of bot \<Rightarrow> bot | Up S \<Rightarrow> Up(update S x (aval' a (Up S))))" |
52046
bc01725d7918 replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents: 47818
diff changeset
   177
"AI (c1;;c2) S = AI c2 (AI c1 S)" |
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   178
"AI (IF b THEN c1 ELSE c2) S =
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   179
  AI c1 (bfilter b True S) \<squnion> AI c2 (bfilter b False S)" |
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   180
"AI (WHILE b DO c) S =
44944
f136409c2cef tuned post fixpoint setup
nipkow
parents: 44932
diff changeset
   181
  bfilter b False (pfp (\<lambda>S. AI c (bfilter b True S)) S)"
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   182
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   183
lemma AI_sound: "(c,s) \<Rightarrow> t \<Longrightarrow> s <:: S \<Longrightarrow> t <:: AI c S"
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44944
diff changeset
   184
proof(induction c arbitrary: s t S)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44656
diff changeset
   185
  case SKIP thus ?case by fastforce
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   186
next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   187
  case Assign thus ?case
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   188
    by (auto simp: lookup_update aval'_sound)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   189
next
47818
151d137f1095 renamed Semi to Seq
nipkow
parents: 45200
diff changeset
   190
  case Seq thus ?case by fastforce
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   191
next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   192
  case If thus ?case by (auto simp: in_rep_join_UpI bfilter_sound)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   193
next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   194
  case (While b c)
44944
f136409c2cef tuned post fixpoint setup
nipkow
parents: 44932
diff changeset
   195
  let ?P = "pfp (\<lambda>S. AI c (bfilter b True S)) S"
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   196
  { fix s t
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   197
    have "(WHILE b DO c,s) \<Rightarrow> t \<Longrightarrow> s <:: ?P \<Longrightarrow>
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   198
          t <:: bfilter b False ?P"
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44944
diff changeset
   199
    proof(induction "WHILE b DO c" s t rule: big_step_induct)
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   200
      case WhileFalse thus ?case by(metis bfilter_sound)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   201
    next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   202
      case WhileTrue show ?case
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   203
        by(rule WhileTrue, rule in_rep_up_trans[OF _ pfp],
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44944
diff changeset
   204
           rule While.IH[OF WhileTrue(2)],
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   205
           rule bfilter_sound[OF WhileTrue.prems], simp add: WhileTrue(1))
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   206
    qed
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   207
  }
44932
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents: 44923
diff changeset
   208
  with in_rep_up_trans[OF `s <:: S` above] While(2,3) AI.simps(5)
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   209
  show ?case by simp
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   210
qed
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   211
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   212
end
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   213
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   214
end