src/HOL/IMP/Abs_Int_ITP/Abs_Int0_ITP.thy
author nipkow
Fri, 10 Aug 2012 17:17:05 +0200
changeset 48759 ff570720ba1c
parent 48480 cb03acfae211
child 51625 bd3358aac5d2
permissions -rw-r--r--
Improved complete lattice formalisation - no more index set.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
     1
(* Author: Tobias Nipkow *)
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
     2
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents: 47602
diff changeset
     3
header "Abstract Interpretation (ITP)"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents: 47602
diff changeset
     4
47602
3d44790b5ab0 reorganised IMP
nipkow
parents: 46334
diff changeset
     5
theory Abs_Int0_ITP
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
     6
imports "~~/src/HOL/ex/Interpretation_with_Defs"
45127
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
     7
        "~~/src/HOL/Library/While_Combinator"
48759
ff570720ba1c Improved complete lattice formalisation - no more index set.
nipkow
parents: 48480
diff changeset
     8
        "Collecting_ITP"
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
     9
begin
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    10
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    11
subsection "Orderings"
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    12
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    13
class preord =
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    14
fixes le :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infix "\<sqsubseteq>" 50)
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    15
assumes le_refl[simp]: "x \<sqsubseteq> x"
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    16
and le_trans: "x \<sqsubseteq> y \<Longrightarrow> y \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> z"
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    17
begin
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    18
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    19
definition mono where "mono f = (\<forall>x y. x \<sqsubseteq> y \<longrightarrow> f x \<sqsubseteq> f y)"
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    20
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    21
lemma monoD: "mono f \<Longrightarrow> x \<sqsubseteq> y \<Longrightarrow> f x \<sqsubseteq> f y" by(simp add: mono_def)
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    22
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    23
lemma mono_comp: "mono f \<Longrightarrow> mono g \<Longrightarrow> mono (g o f)"
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    24
by(simp add: mono_def)
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    25
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    26
declare le_trans[trans]
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    27
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    28
end
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    29
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    30
text{* Note: no antisymmetry. Allows implementations where some abstract
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    31
element is implemented by two different values @{prop "x \<noteq> y"}
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    32
such that @{prop"x \<sqsubseteq> y"} and @{prop"y \<sqsubseteq> x"}. Antisymmetry is not
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    33
needed because we never compare elements for equality but only for @{text"\<sqsubseteq>"}.
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    34
*}
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    35
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    36
class SL_top = preord +
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    37
fixes join :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<squnion>" 65)
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    38
fixes Top :: "'a" ("\<top>")
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    39
assumes join_ge1 [simp]: "x \<sqsubseteq> x \<squnion> y"
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    40
and join_ge2 [simp]: "y \<sqsubseteq> x \<squnion> y"
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    41
and join_least: "x \<sqsubseteq> z \<Longrightarrow> y \<sqsubseteq> z \<Longrightarrow> x \<squnion> y \<sqsubseteq> z"
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    42
and top[simp]: "x \<sqsubseteq> \<top>"
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    43
begin
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    44
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    45
lemma join_le_iff[simp]: "x \<squnion> y \<sqsubseteq> z \<longleftrightarrow> x \<sqsubseteq> z \<and> y \<sqsubseteq> z"
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    46
by (metis join_ge1 join_ge2 join_least le_trans)
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    47
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    48
lemma le_join_disj: "x \<sqsubseteq> y \<or> x \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> y \<squnion> z"
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    49
by (metis join_ge1 join_ge2 le_trans)
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    50
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    51
end
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    52
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    53
instantiation "fun" :: (type, SL_top) SL_top
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    54
begin
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    55
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    56
definition "f \<sqsubseteq> g = (ALL x. f x \<sqsubseteq> g x)"
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    57
definition "f \<squnion> g = (\<lambda>x. f x \<squnion> g x)"
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    58
definition "\<top> = (\<lambda>x. \<top>)"
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    59
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    60
lemma join_apply[simp]: "(f \<squnion> g) x = f x \<squnion> g x"
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    61
by (simp add: join_fun_def)
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    62
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    63
instance
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    64
proof
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    65
  case goal2 thus ?case by (metis le_fun_def preord_class.le_trans)
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    66
qed (simp_all add: le_fun_def Top_fun_def)
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    67
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    68
end
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    69
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    70
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    71
instantiation acom :: (preord) preord
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    72
begin
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    73
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    74
fun le_acom :: "('a::preord)acom \<Rightarrow> 'a acom \<Rightarrow> bool" where
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    75
"le_acom (SKIP {S}) (SKIP {S'}) = (S \<sqsubseteq> S')" |
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    76
"le_acom (x ::= e {S}) (x' ::= e' {S'}) = (x=x' \<and> e=e' \<and> S \<sqsubseteq> S')" |
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    77
"le_acom (c1;c2) (c1';c2') = (le_acom c1 c1' \<and> le_acom c2 c2')" |
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    78
"le_acom (IF b THEN c1 ELSE c2 {S}) (IF b' THEN c1' ELSE c2' {S'}) =
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    79
  (b=b' \<and> le_acom c1 c1' \<and> le_acom c2 c2' \<and> S \<sqsubseteq> S')" |
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    80
"le_acom ({Inv} WHILE b DO c {P}) ({Inv'} WHILE b' DO c' {P'}) =
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    81
  (b=b' \<and> le_acom c c' \<and> Inv \<sqsubseteq> Inv' \<and> P \<sqsubseteq> P')" |
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    82
"le_acom _ _ = False"
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    83
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    84
lemma [simp]: "SKIP {S} \<sqsubseteq> c \<longleftrightarrow> (\<exists>S'. c = SKIP {S'} \<and> S \<sqsubseteq> S')"
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    85
by (cases c) auto
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    86
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    87
lemma [simp]: "x ::= e {S} \<sqsubseteq> c \<longleftrightarrow> (\<exists>S'. c = x ::= e {S'} \<and> S \<sqsubseteq> S')"
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    88
by (cases c) auto
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    89
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    90
lemma [simp]: "c1;c2 \<sqsubseteq> c \<longleftrightarrow> (\<exists>c1' c2'. c = c1';c2' \<and> c1 \<sqsubseteq> c1' \<and> c2 \<sqsubseteq> c2')"
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    91
by (cases c) auto
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    92
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    93
lemma [simp]: "IF b THEN c1 ELSE c2 {S} \<sqsubseteq> c \<longleftrightarrow>
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    94
  (\<exists>c1' c2' S'. c = IF b THEN c1' ELSE c2' {S'} \<and> c1 \<sqsubseteq> c1' \<and> c2 \<sqsubseteq> c2' \<and> S \<sqsubseteq> S')"
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    95
by (cases c) auto
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    96
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    97
lemma [simp]: "{Inv} WHILE b DO c {P} \<sqsubseteq> w \<longleftrightarrow>
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    98
  (\<exists>Inv' c' P'. w = {Inv'} WHILE b DO c' {P'} \<and> c \<sqsubseteq> c' \<and> Inv \<sqsubseteq> Inv' \<and> P \<sqsubseteq> P')"
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    99
by (cases w) auto
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   100
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   101
instance
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   102
proof
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   103
  case goal1 thus ?case by (induct x) auto
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   104
next
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   105
  case goal2 thus ?case
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   106
  apply(induct x y arbitrary: z rule: le_acom.induct)
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   107
  apply (auto intro: le_trans)
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   108
  done
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   109
qed
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   110
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   111
end
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   112
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   113
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   114
subsubsection "Lifting"
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   115
46078
nipkow
parents: 46070
diff changeset
   116
instantiation option :: (preord)preord
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   117
begin
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   118
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   119
fun le_option where
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   120
"Some x \<sqsubseteq> Some y = (x \<sqsubseteq> y)" |
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   121
"None \<sqsubseteq> y = True" |
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   122
"Some _ \<sqsubseteq> None = False"
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   123
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   124
lemma [simp]: "(x \<sqsubseteq> None) = (x = None)"
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   125
by (cases x) simp_all
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   126
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   127
lemma [simp]: "(Some x \<sqsubseteq> u) = (\<exists>y. u = Some y & x \<sqsubseteq> y)"
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   128
by (cases u) auto
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   129
46078
nipkow
parents: 46070
diff changeset
   130
instance proof
nipkow
parents: 46070
diff changeset
   131
  case goal1 show ?case by(cases x, simp_all)
nipkow
parents: 46070
diff changeset
   132
next
nipkow
parents: 46070
diff changeset
   133
  case goal2 thus ?case
nipkow
parents: 46070
diff changeset
   134
    by(cases z, simp, cases y, simp, cases x, auto intro: le_trans)
nipkow
parents: 46070
diff changeset
   135
qed
nipkow
parents: 46070
diff changeset
   136
nipkow
parents: 46070
diff changeset
   137
end
nipkow
parents: 46070
diff changeset
   138
nipkow
parents: 46070
diff changeset
   139
instantiation option :: (SL_top)SL_top
nipkow
parents: 46070
diff changeset
   140
begin
nipkow
parents: 46070
diff changeset
   141
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   142
fun join_option where
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   143
"Some x \<squnion> Some y = Some(x \<squnion> y)" |
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   144
"None \<squnion> y = y" |
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   145
"x \<squnion> None = x"
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   146
46039
nipkow
parents: 45903
diff changeset
   147
lemma join_None2[simp]: "x \<squnion> None = x"
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   148
by (cases x) simp_all
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   149
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   150
definition "\<top> = Some \<top>"
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   151
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   152
instance proof
46078
nipkow
parents: 46070
diff changeset
   153
  case goal1 thus ?case by(cases x, simp, cases y, simp_all)
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   154
next
46078
nipkow
parents: 46070
diff changeset
   155
  case goal2 thus ?case by(cases y, simp, cases x, simp_all)
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   156
next
46078
nipkow
parents: 46070
diff changeset
   157
  case goal3 thus ?case by(cases z, simp, cases y, simp, cases x, simp_all)
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   158
next
46078
nipkow
parents: 46070
diff changeset
   159
  case goal4 thus ?case by(cases x, simp_all add: Top_option_def)
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   160
qed
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   161
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   162
end
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   163
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   164
definition bot_acom :: "com \<Rightarrow> ('a::SL_top)option acom" ("\<bottom>\<^sub>c") where
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   165
"\<bottom>\<^sub>c = anno None"
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   166
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   167
lemma strip_bot_acom[simp]: "strip(\<bottom>\<^sub>c c) = c"
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   168
by(simp add: bot_acom_def)
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   169
45127
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   170
lemma bot_acom[rule_format]: "strip c' = c \<longrightarrow> \<bottom>\<^sub>c c \<sqsubseteq> c'"
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   171
apply(induct c arbitrary: c')
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   172
apply (simp_all add: bot_acom_def)
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   173
 apply(induct_tac c')
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   174
  apply simp_all
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   175
 apply(induct_tac c')
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   176
  apply simp_all
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   177
 apply(induct_tac c')
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   178
  apply simp_all
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   179
 apply(induct_tac c')
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   180
  apply simp_all
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   181
 apply(induct_tac c')
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   182
  apply simp_all
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   183
done
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   184
45127
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   185
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   186
subsubsection "Post-fixed point iteration"
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   187
45127
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   188
definition
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   189
  pfp :: "(('a::preord) \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a option" where
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   190
"pfp f = while_option (\<lambda>x. \<not> f x \<sqsubseteq> x) f"
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   191
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   192
lemma pfp_pfp: assumes "pfp f x0 = Some x" shows "f x \<sqsubseteq> x"
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   193
using while_option_stop[OF assms[simplified pfp_def]] by simp
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   194
45127
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   195
lemma pfp_least:
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   196
assumes mono: "\<And>x y. x \<sqsubseteq> y \<Longrightarrow> f x \<sqsubseteq> f y"
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   197
and "f p \<sqsubseteq> p" and "x0 \<sqsubseteq> p" and "pfp f x0 = Some x" shows "x \<sqsubseteq> p"
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   198
proof-
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   199
  { fix x assume "x \<sqsubseteq> p"
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   200
    hence  "f x \<sqsubseteq> f p" by(rule mono)
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   201
    from this `f p \<sqsubseteq> p` have "f x \<sqsubseteq> p" by(rule le_trans)
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   202
  }
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   203
  thus "x \<sqsubseteq> p" using assms(2-) while_option_rule[where P = "%x. x \<sqsubseteq> p"]
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   204
    unfolding pfp_def by blast
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   205
qed
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   206
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   207
definition
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   208
 lpfp\<^isub>c :: "(('a::SL_top)option acom \<Rightarrow> 'a option acom) \<Rightarrow> com \<Rightarrow> 'a option acom option" where
45127
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   209
"lpfp\<^isub>c f c = pfp f (\<bottom>\<^sub>c c)"
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   210
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   211
lemma lpfpc_pfp: "lpfp\<^isub>c f c0 = Some c \<Longrightarrow> f c \<sqsubseteq> c"
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   212
by(simp add: pfp_pfp lpfp\<^isub>c_def)
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   213
45127
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   214
lemma strip_pfp:
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   215
assumes "\<And>x. g(f x) = g x" and "pfp f x0 = Some x" shows "g x = g x0"
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   216
using assms while_option_rule[where P = "%x. g x = g x0" and c = f]
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   217
unfolding pfp_def by metis
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   218
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   219
lemma strip_lpfpc: assumes "\<And>c. strip(f c) = strip c" and "lpfp\<^isub>c f c = Some c'"
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   220
shows "strip c' = c"
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   221
using assms(1) strip_pfp[OF _ assms(2)[simplified lpfp\<^isub>c_def]]
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   222
by(metis strip_bot_acom)
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   223
45127
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   224
lemma lpfpc_least:
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   225
assumes mono: "\<And>x y. x \<sqsubseteq> y \<Longrightarrow> f x \<sqsubseteq> f y"
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   226
and "strip p = c0" and "f p \<sqsubseteq> p" and lp: "lpfp\<^isub>c f c0 = Some c" shows "c \<sqsubseteq> p"
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   227
using pfp_least[OF _ _ bot_acom[OF `strip p = c0`] lp[simplified lpfp\<^isub>c_def]]
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   228
  mono `f p \<sqsubseteq> p`
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   229
by blast
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   230
45127
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   231
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   232
subsection "Abstract Interpretation"
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   233
46039
nipkow
parents: 45903
diff changeset
   234
definition \<gamma>_fun :: "('a \<Rightarrow> 'b set) \<Rightarrow> ('c \<Rightarrow> 'a) \<Rightarrow> ('c \<Rightarrow> 'b)set" where
nipkow
parents: 45903
diff changeset
   235
"\<gamma>_fun \<gamma> F = {f. \<forall>x. f x \<in> \<gamma>(F x)}"
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   236
46039
nipkow
parents: 45903
diff changeset
   237
fun \<gamma>_option :: "('a \<Rightarrow> 'b set) \<Rightarrow> 'a option \<Rightarrow> 'b set" where
nipkow
parents: 45903
diff changeset
   238
"\<gamma>_option \<gamma> None = {}" |
nipkow
parents: 45903
diff changeset
   239
"\<gamma>_option \<gamma> (Some a) = \<gamma> a"
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   240
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   241
text{* The interface for abstract values: *}
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   242
45127
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   243
locale Val_abs =
46063
81ebd0cdb300 tuned types
nipkow
parents: 46039
diff changeset
   244
fixes \<gamma> :: "'av::SL_top \<Rightarrow> val set"
46039
nipkow
parents: 45903
diff changeset
   245
  assumes mono_gamma: "a \<sqsubseteq> b \<Longrightarrow> \<gamma> a \<subseteq> \<gamma> b"
nipkow
parents: 45903
diff changeset
   246
  and gamma_Top[simp]: "\<gamma> \<top> = UNIV"
46063
81ebd0cdb300 tuned types
nipkow
parents: 46039
diff changeset
   247
fixes num' :: "val \<Rightarrow> 'av"
81ebd0cdb300 tuned types
nipkow
parents: 46039
diff changeset
   248
and plus' :: "'av \<Rightarrow> 'av \<Rightarrow> 'av"
46039
nipkow
parents: 45903
diff changeset
   249
  assumes gamma_num': "n : \<gamma>(num' n)"
nipkow
parents: 45903
diff changeset
   250
  and gamma_plus':
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   251
 "n1 : \<gamma> a1 \<Longrightarrow> n2 : \<gamma> a2 \<Longrightarrow> n1+n2 : \<gamma>(plus' a1 a2)"
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   252
46063
81ebd0cdb300 tuned types
nipkow
parents: 46039
diff changeset
   253
type_synonym 'av st = "(vname \<Rightarrow> 'av)"
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   254
46063
81ebd0cdb300 tuned types
nipkow
parents: 46039
diff changeset
   255
locale Abs_Int_Fun = Val_abs \<gamma> for \<gamma> :: "'av::SL_top \<Rightarrow> val set"
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   256
begin
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   257
46063
81ebd0cdb300 tuned types
nipkow
parents: 46039
diff changeset
   258
fun aval' :: "aexp \<Rightarrow> 'av st \<Rightarrow> 'av" where
46039
nipkow
parents: 45903
diff changeset
   259
"aval' (N n) S = num' n" |
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   260
"aval' (V x) S = S x" |
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   261
"aval' (Plus a1 a2) S = plus' (aval' a1 S) (aval' a2 S)"
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   262
46063
81ebd0cdb300 tuned types
nipkow
parents: 46039
diff changeset
   263
fun step' :: "'av st option \<Rightarrow> 'av st option acom \<Rightarrow> 'av st option acom"
45127
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   264
 where
45655
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
   265
"step' S (SKIP {P}) = (SKIP {S})" |
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
   266
"step' S (x ::= e {P}) =
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   267
  x ::= e {case S of None \<Rightarrow> None | Some S \<Rightarrow> Some(S(x := aval' e S))}" |
45655
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
   268
"step' S (c1; c2) = step' S c1; step' (post c1) c2" |
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
   269
"step' S (IF b THEN c1 ELSE c2 {P}) =
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
   270
   IF b THEN step' S c1 ELSE step' S c2 {post c1 \<squnion> post c2}" |
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
   271
"step' S ({Inv} WHILE b DO c {P}) =
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
   272
  {S \<squnion> post c} WHILE b DO (step' Inv c) {Inv}"
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   273
46063
81ebd0cdb300 tuned types
nipkow
parents: 46039
diff changeset
   274
definition AI :: "com \<Rightarrow> 'av st option acom option" where
45655
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
   275
"AI = lpfp\<^isub>c (step' \<top>)"
45127
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   276
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   277
45655
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
   278
lemma strip_step'[simp]: "strip(step' S c) = strip c"
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   279
by(induct c arbitrary: S) (simp_all add: Let_def)
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   280
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   281
46063
81ebd0cdb300 tuned types
nipkow
parents: 46039
diff changeset
   282
abbreviation \<gamma>\<^isub>f :: "'av st \<Rightarrow> state set"
46039
nipkow
parents: 45903
diff changeset
   283
where "\<gamma>\<^isub>f == \<gamma>_fun \<gamma>"
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   284
46063
81ebd0cdb300 tuned types
nipkow
parents: 46039
diff changeset
   285
abbreviation \<gamma>\<^isub>o :: "'av st option \<Rightarrow> state set"
46039
nipkow
parents: 45903
diff changeset
   286
where "\<gamma>\<^isub>o == \<gamma>_option \<gamma>\<^isub>f"
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   287
46063
81ebd0cdb300 tuned types
nipkow
parents: 46039
diff changeset
   288
abbreviation \<gamma>\<^isub>c :: "'av st option acom \<Rightarrow> state set acom"
46039
nipkow
parents: 45903
diff changeset
   289
where "\<gamma>\<^isub>c == map_acom \<gamma>\<^isub>o"
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   290
46039
nipkow
parents: 45903
diff changeset
   291
lemma gamma_f_Top[simp]: "\<gamma>\<^isub>f Top = UNIV"
nipkow
parents: 45903
diff changeset
   292
by(simp add: Top_fun_def \<gamma>_fun_def)
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   293
46039
nipkow
parents: 45903
diff changeset
   294
lemma gamma_o_Top[simp]: "\<gamma>\<^isub>o Top = UNIV"
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   295
by (simp add: Top_option_def)
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   296
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   297
(* FIXME (maybe also le \<rightarrow> sqle?) *)
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   298
46039
nipkow
parents: 45903
diff changeset
   299
lemma mono_gamma_f: "f \<sqsubseteq> g \<Longrightarrow> \<gamma>\<^isub>f f \<subseteq> \<gamma>\<^isub>f g"
nipkow
parents: 45903
diff changeset
   300
by(auto simp: le_fun_def \<gamma>_fun_def dest: mono_gamma)
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   301
46039
nipkow
parents: 45903
diff changeset
   302
lemma mono_gamma_o:
nipkow
parents: 45903
diff changeset
   303
  "sa \<sqsubseteq> sa' \<Longrightarrow> \<gamma>\<^isub>o sa \<subseteq> \<gamma>\<^isub>o sa'"
nipkow
parents: 45903
diff changeset
   304
by(induction sa sa' rule: le_option.induct)(simp_all add: mono_gamma_f)
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   305
46039
nipkow
parents: 45903
diff changeset
   306
lemma mono_gamma_c: "ca \<sqsubseteq> ca' \<Longrightarrow> \<gamma>\<^isub>c ca \<le> \<gamma>\<^isub>c ca'"
nipkow
parents: 45903
diff changeset
   307
by (induction ca ca' rule: le_acom.induct) (simp_all add:mono_gamma_o)
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   308
45127
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   309
text{* Soundness: *}
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   310
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   311
lemma aval'_sound: "s : \<gamma>\<^isub>f S \<Longrightarrow> aval a s : \<gamma>(aval' a S)"
46039
nipkow
parents: 45903
diff changeset
   312
by (induct a) (auto simp: gamma_num' gamma_plus' \<gamma>_fun_def)
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   313
46039
nipkow
parents: 45903
diff changeset
   314
lemma in_gamma_update:
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   315
  "\<lbrakk> s : \<gamma>\<^isub>f S; i : \<gamma> a \<rbrakk> \<Longrightarrow> s(x := i) : \<gamma>\<^isub>f(S(x := a))"
46039
nipkow
parents: 45903
diff changeset
   316
by(simp add: \<gamma>_fun_def)
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   317
46068
b9d4ec0f79ac tuned proofs
nipkow
parents: 46067
diff changeset
   318
lemma step_preserves_le:
46334
nipkow
parents: 46225
diff changeset
   319
  "\<lbrakk> S \<subseteq> \<gamma>\<^isub>o S'; c \<le> \<gamma>\<^isub>c c' \<rbrakk> \<Longrightarrow> step S c \<le> \<gamma>\<^isub>c (step' S' c')"
nipkow
parents: 46225
diff changeset
   320
proof(induction c arbitrary: c' S S')
46068
b9d4ec0f79ac tuned proofs
nipkow
parents: 46067
diff changeset
   321
  case SKIP thus ?case by(auto simp:SKIP_le map_acom_SKIP)
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   322
next
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   323
  case Assign thus ?case
46068
b9d4ec0f79ac tuned proofs
nipkow
parents: 46067
diff changeset
   324
    by (fastforce simp: Assign_le  map_acom_Assign intro: aval'_sound in_gamma_update
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   325
      split: option.splits del:subsetD)
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   326
next
47818
151d137f1095 renamed Semi to Seq
nipkow
parents: 47613
diff changeset
   327
  case Seq thus ?case apply (auto simp: Seq_le map_acom_Seq)
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   328
    by (metis le_post post_map_acom)
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   329
next
46334
nipkow
parents: 46225
diff changeset
   330
  case (If b c1 c2 P)
nipkow
parents: 46225
diff changeset
   331
  then obtain c1' c2' P' where
nipkow
parents: 46225
diff changeset
   332
      "c' = IF b THEN c1' ELSE c2' {P'}"
nipkow
parents: 46225
diff changeset
   333
      "P \<subseteq> \<gamma>\<^isub>o P'" "c1 \<le> \<gamma>\<^isub>c c1'" "c2 \<le> \<gamma>\<^isub>c c2'"
46068
b9d4ec0f79ac tuned proofs
nipkow
parents: 46067
diff changeset
   334
    by (fastforce simp: If_le map_acom_If)
46334
nipkow
parents: 46225
diff changeset
   335
  moreover have "post c1 \<subseteq> \<gamma>\<^isub>o(post c1' \<squnion> post c2')"
nipkow
parents: 46225
diff changeset
   336
    by (metis (no_types) `c1 \<le> \<gamma>\<^isub>c c1'` join_ge1 le_post mono_gamma_o order_trans post_map_acom)
nipkow
parents: 46225
diff changeset
   337
  moreover have "post c2 \<subseteq> \<gamma>\<^isub>o(post c1' \<squnion> post c2')"
nipkow
parents: 46225
diff changeset
   338
    by (metis (no_types) `c2 \<le> \<gamma>\<^isub>c c2'` join_ge2 le_post mono_gamma_o order_trans post_map_acom)
46067
a03bf644cb27 tuned var names
nipkow
parents: 46066
diff changeset
   339
  ultimately show ?case using `S \<subseteq> \<gamma>\<^isub>o S'` by (simp add: If.IH subset_iff)
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   340
next
46334
nipkow
parents: 46225
diff changeset
   341
  case (While I b c1 P)
nipkow
parents: 46225
diff changeset
   342
  then obtain c1' I' P' where
nipkow
parents: 46225
diff changeset
   343
    "c' = {I'} WHILE b DO c1' {P'}"
nipkow
parents: 46225
diff changeset
   344
    "I \<subseteq> \<gamma>\<^isub>o I'" "P \<subseteq> \<gamma>\<^isub>o P'" "c1 \<le> \<gamma>\<^isub>c c1'"
46068
b9d4ec0f79ac tuned proofs
nipkow
parents: 46067
diff changeset
   345
    by (fastforce simp: map_acom_While While_le)
46334
nipkow
parents: 46225
diff changeset
   346
  moreover have "S \<union> post c1 \<subseteq> \<gamma>\<^isub>o (S' \<squnion> post c1')"
nipkow
parents: 46225
diff changeset
   347
    using `S \<subseteq> \<gamma>\<^isub>o S'` le_post[OF `c1 \<le> \<gamma>\<^isub>c c1'`, simplified]
46039
nipkow
parents: 45903
diff changeset
   348
    by (metis (no_types) join_ge1 join_ge2 le_sup_iff mono_gamma_o order_trans)
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   349
  ultimately show ?case by (simp add: While.IH subset_iff)
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   350
qed
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   351
46070
nipkow
parents: 46068
diff changeset
   352
lemma AI_sound: "AI c = Some c' \<Longrightarrow> CS c \<le> \<gamma>\<^isub>c c'"
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   353
proof(simp add: CS_def AI_def)
45655
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
   354
  assume 1: "lpfp\<^isub>c (step' \<top>) c = Some c'"
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
   355
  have 2: "step' \<top> c' \<sqsubseteq> c'" by(rule lpfpc_pfp[OF 1])
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
   356
  have 3: "strip (\<gamma>\<^isub>c (step' \<top> c')) = c"
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   357
    by(simp add: strip_lpfpc[OF _ 1])
46066
e81411bfa7ef tuned argument order
nipkow
parents: 46063
diff changeset
   358
  have "lfp (step UNIV) c \<le> \<gamma>\<^isub>c (step' \<top> c')"
45903
02dd9319dcb7 improved indexed complete lattice
nipkow
parents: 45655
diff changeset
   359
  proof(rule lfp_lowerbound[simplified,OF 3])
45655
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
   360
    show "step UNIV (\<gamma>\<^isub>c (step' \<top> c')) \<le> \<gamma>\<^isub>c (step' \<top> c')"
46068
b9d4ec0f79ac tuned proofs
nipkow
parents: 46067
diff changeset
   361
    proof(rule step_preserves_le[OF _ _])
46039
nipkow
parents: 45903
diff changeset
   362
      show "UNIV \<subseteq> \<gamma>\<^isub>o \<top>" by simp
nipkow
parents: 45903
diff changeset
   363
      show "\<gamma>\<^isub>c (step' \<top> c') \<le> \<gamma>\<^isub>c c'" by(rule mono_gamma_c[OF 2])
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   364
    qed
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   365
  qed
46068
b9d4ec0f79ac tuned proofs
nipkow
parents: 46067
diff changeset
   366
  with 2 show "lfp (step UNIV) c \<le> \<gamma>\<^isub>c c'"
46039
nipkow
parents: 45903
diff changeset
   367
    by (blast intro: mono_gamma_c order_trans)
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   368
qed
45127
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   369
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   370
end
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   371
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   372
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   373
subsubsection "Monotonicity"
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   374
46153
nipkow
parents: 46078
diff changeset
   375
lemma mono_post: "c \<sqsubseteq> c' \<Longrightarrow> post c \<sqsubseteq> post c'"
nipkow
parents: 46078
diff changeset
   376
by(induction c c' rule: le_acom.induct) (auto)
nipkow
parents: 46078
diff changeset
   377
45127
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   378
locale Abs_Int_Fun_mono = Abs_Int_Fun +
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   379
assumes mono_plus': "a1 \<sqsubseteq> b1 \<Longrightarrow> a2 \<sqsubseteq> b2 \<Longrightarrow> plus' a1 a2 \<sqsubseteq> plus' b1 b2"
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   380
begin
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   381
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   382
lemma mono_aval': "S \<sqsubseteq> S' \<Longrightarrow> aval' e S \<sqsubseteq> aval' e S'"
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   383
by(induction e)(auto simp: le_fun_def mono_plus')
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   384
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   385
lemma mono_update: "a \<sqsubseteq> a' \<Longrightarrow> S \<sqsubseteq> S' \<Longrightarrow> S(x := a) \<sqsubseteq> S'(x := a')"
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   386
by(simp add: le_fun_def)
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   387
46153
nipkow
parents: 46078
diff changeset
   388
lemma mono_step': "S \<sqsubseteq> S' \<Longrightarrow> c \<sqsubseteq> c' \<Longrightarrow> step' S c \<sqsubseteq> step' S' c'"
nipkow
parents: 46078
diff changeset
   389
apply(induction c c' arbitrary: S S' rule: le_acom.induct)
nipkow
parents: 46078
diff changeset
   390
apply (auto simp: Let_def mono_update mono_aval' mono_post le_join_disj
nipkow
parents: 46078
diff changeset
   391
            split: option.split)
45127
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   392
done
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   393
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   394
end
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   395
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   396
text{* Problem: not executable because of the comparison of abstract states,
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   397
i.e. functions, in the post-fixedpoint computation. *}
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   398
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   399
end