src/HOL/IMP/Abs_Int0_fun.thy
author nipkow
Tue, 13 Dec 2011 21:15:38 +0100
changeset 45840 dadd139192d1
parent 45655 a49f9428aba4
child 45903 02dd9319dcb7
permissions -rw-r--r--
added concrete syntax
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
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
     3
header "Abstract Interpretation"
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
     4
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
     5
theory Abs_Int0_fun
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"
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
     8
        Collecting
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
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   116
instantiation option :: (SL_top)SL_top
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
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   130
fun join_option where
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   131
"Some x \<squnion> Some y = Some(x \<squnion> y)" |
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   132
"None \<squnion> y = y" |
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   133
"x \<squnion> None = x"
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   134
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   135
lemma [simp]: "x \<squnion> None = x"
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   136
by (cases x) simp_all
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   137
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   138
definition "\<top> = Some \<top>"
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   139
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   140
instance proof
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   141
  case goal1 show ?case by(cases x, simp_all)
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   142
next
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   143
  case goal2 thus ?case
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   144
    by(cases z, simp, cases y, simp, cases x, auto intro: le_trans)
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   145
next
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   146
  case goal3 thus ?case by(cases x, simp, cases y, simp_all)
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   147
next
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   148
  case goal4 thus ?case by(cases y, simp, cases x, simp_all)
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   149
next
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   150
  case goal5 thus ?case by(cases z, simp, cases y, simp, cases x, simp_all)
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   151
next
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   152
  case goal6 thus ?case by(cases x, simp_all add: Top_option_def)
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   153
qed
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   154
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   155
end
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   156
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   157
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
   158
"\<bottom>\<^sub>c = anno None"
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   159
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   160
lemma strip_bot_acom[simp]: "strip(\<bottom>\<^sub>c c) = c"
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   161
by(simp add: bot_acom_def)
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   162
45127
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   163
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
   164
apply(induct c arbitrary: c')
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   165
apply (simp_all add: bot_acom_def)
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   166
 apply(induct_tac c')
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   167
  apply simp_all
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   168
 apply(induct_tac c')
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   169
  apply simp_all
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   170
 apply(induct_tac c')
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   171
  apply simp_all
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   172
 apply(induct_tac c')
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   173
  apply simp_all
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   174
 apply(induct_tac c')
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   175
  apply simp_all
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   176
done
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   177
45127
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   178
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   179
subsubsection "Post-fixed point iteration"
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   180
45127
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   181
definition
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   182
  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
   183
"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
   184
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   185
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
   186
using while_option_stop[OF assms[simplified pfp_def]] by simp
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
lemma pfp_least:
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   189
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
   190
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
   191
proof-
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   192
  { fix x assume "x \<sqsubseteq> p"
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   193
    hence  "f x \<sqsubseteq> f p" by(rule mono)
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   194
    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
   195
  }
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   196
  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
   197
    unfolding pfp_def by blast
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   198
qed
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   199
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   200
definition
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   201
 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
   202
"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
   203
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   204
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
   205
by(simp add: pfp_pfp lpfp\<^isub>c_def)
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   206
45127
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   207
lemma strip_pfp:
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   208
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
   209
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
   210
unfolding pfp_def by metis
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   211
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   212
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
   213
shows "strip c' = c"
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   214
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
   215
by(metis strip_bot_acom)
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   216
45127
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   217
lemma lpfpc_least:
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   218
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
   219
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
   220
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
   221
  mono `f p \<sqsubseteq> p`
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   222
by blast
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
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   225
subsection "Abstract Interpretation"
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   226
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   227
definition rep_fun :: "('a \<Rightarrow> 'b set) \<Rightarrow> ('c \<Rightarrow> 'a) \<Rightarrow> ('c \<Rightarrow> 'b)set" where
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   228
"rep_fun rep F = {f. \<forall>x. f x \<in> rep(F x)}"
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   229
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   230
fun rep_option :: "('a \<Rightarrow> 'b set) \<Rightarrow> 'a option \<Rightarrow> 'b set" where
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   231
"rep_option rep None = {}" |
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   232
"rep_option rep (Some a) = rep a"
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   233
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   234
text{* The interface for abstract values: *}
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   235
45127
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   236
locale Val_abs =
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   237
fixes rep :: "'a::SL_top \<Rightarrow> val set" ("\<gamma>")
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   238
  assumes mono_rep: "a \<sqsubseteq> b \<Longrightarrow> \<gamma> a \<subseteq> \<gamma> b"
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   239
  and rep_Top[simp]: "\<gamma> \<top> = UNIV"
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   240
fixes num' :: "val \<Rightarrow> 'a"
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   241
and plus' :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   242
  assumes rep_num': "n : \<gamma>(num' n)"
45127
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   243
  and rep_plus':
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   244
 "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
   245
45212
e87feee00a4c renamed name -> vname
nipkow
parents: 45127
diff changeset
   246
type_synonym 'a st = "(vname \<Rightarrow> 'a)"
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   247
45127
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   248
locale Abs_Int_Fun = Val_abs
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   249
begin
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   250
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   251
fun aval' :: "aexp \<Rightarrow> 'a st \<Rightarrow> 'a" where
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   252
"aval' (N n) _ = num' n" |
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   253
"aval' (V x) S = S x" |
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   254
"aval' (Plus a1 a2) S = plus' (aval' a1 S) (aval' a2 S)"
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   255
45655
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
   256
fun step' :: "'a st option \<Rightarrow> 'a st option acom \<Rightarrow> 'a st option acom"
45127
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   257
 where
45655
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
   258
"step' S (SKIP {P}) = (SKIP {S})" |
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
   259
"step' S (x ::= e {P}) =
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   260
  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
   261
"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
   262
"step' S (IF b THEN c1 ELSE c2 {P}) =
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
   263
   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
   264
"step' S ({Inv} WHILE b DO c {P}) =
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
   265
  {S \<squnion> post c} WHILE b DO (step' Inv c) {Inv}"
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   266
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   267
definition AI :: "com \<Rightarrow> 'a st option acom option" where
45655
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
   268
"AI = lpfp\<^isub>c (step' \<top>)"
45127
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   269
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   270
45655
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
   271
lemma strip_step'[simp]: "strip(step' S c) = strip c"
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   272
by(induct c arbitrary: S) (simp_all add: Let_def)
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   273
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   274
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   275
abbreviation rep_f :: "'a st \<Rightarrow> state set" ("\<gamma>\<^isub>f")
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   276
where "\<gamma>\<^isub>f == rep_fun \<gamma>"
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   277
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   278
abbreviation rep_u :: "'a st option \<Rightarrow> state set" ("\<gamma>\<^isub>u")
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   279
where "\<gamma>\<^isub>u == rep_option \<gamma>\<^isub>f"
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   280
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   281
abbreviation rep_c :: "'a st option acom \<Rightarrow> state set acom" ("\<gamma>\<^isub>c")
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   282
where "\<gamma>\<^isub>c == map_acom \<gamma>\<^isub>u"
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   283
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   284
lemma rep_f_Top[simp]: "\<gamma>\<^isub>f Top = UNIV"
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   285
by(simp add: Top_fun_def rep_fun_def)
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   286
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   287
lemma rep_u_Top[simp]: "\<gamma>\<^isub>u Top = UNIV"
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   288
by (simp add: Top_option_def)
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   289
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   290
(* FIXME (maybe also le \<rightarrow> sqle?) *)
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   291
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   292
lemma mono_rep_f: "f \<sqsubseteq> g \<Longrightarrow> \<gamma>\<^isub>f f \<subseteq> \<gamma>\<^isub>f g"
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   293
by(auto simp: le_fun_def rep_fun_def dest: mono_rep)
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   294
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   295
lemma mono_rep_u:
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   296
  "sa \<sqsubseteq> sa' \<Longrightarrow> \<gamma>\<^isub>u sa \<subseteq> \<gamma>\<^isub>u sa'"
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   297
by(induction sa sa' rule: le_option.induct)(simp_all add: mono_rep_f)
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   298
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   299
lemma mono_rep_c: "ca \<sqsubseteq> ca' \<Longrightarrow> \<gamma>\<^isub>c ca \<le> \<gamma>\<^isub>c ca'"
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   300
by (induction ca ca' rule: le_acom.induct) (simp_all add:mono_rep_u)
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   301
45127
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   302
text{* Soundness: *}
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   303
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   304
lemma aval'_sound: "s : \<gamma>\<^isub>f S \<Longrightarrow> aval a s : \<gamma>(aval' a S)"
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   305
by (induct a) (auto simp: rep_num' rep_plus' rep_fun_def)
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   306
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   307
lemma in_rep_update:
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   308
  "\<lbrakk> s : \<gamma>\<^isub>f S; i : \<gamma> a \<rbrakk> \<Longrightarrow> s(x := i) : \<gamma>\<^isub>f(S(x := a))"
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   309
by(simp add: rep_fun_def)
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 step_preserves_le2:
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   312
  "\<lbrakk> S \<subseteq> \<gamma>\<^isub>u sa; cs \<le> \<gamma>\<^isub>c ca; strip cs = c; strip ca = c \<rbrakk>
45655
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
   313
   \<Longrightarrow> step S cs \<le> \<gamma>\<^isub>c (step' sa ca)"
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   314
proof(induction c arbitrary: cs ca S sa)
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   315
  case SKIP thus ?case
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   316
    by(auto simp:strip_eq_SKIP)
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   317
next
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   318
  case Assign thus ?case
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   319
    by (fastforce simp: strip_eq_Assign intro: aval'_sound in_rep_update
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   320
      split: option.splits del:subsetD)
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   321
next
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   322
  case Semi thus ?case apply (auto simp: strip_eq_Semi)
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   323
    by (metis le_post post_map_acom)
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   324
next
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   325
  case (If b c1 c2)
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   326
  then obtain cs1 cs2 ca1 ca2 P Pa where
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   327
      "cs= IF b THEN cs1 ELSE cs2 {P}" "ca= IF b THEN ca1 ELSE ca2 {Pa}"
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   328
      "P \<subseteq> \<gamma>\<^isub>u Pa" "cs1 \<le> \<gamma>\<^isub>c ca1" "cs2 \<le> \<gamma>\<^isub>c ca2"
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   329
      "strip cs1 = c1" "strip ca1 = c1" "strip cs2 = c2" "strip ca2 = c2"
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   330
    by (fastforce simp: strip_eq_If)
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   331
  moreover have "post cs1 \<subseteq> \<gamma>\<^isub>u(post ca1 \<squnion> post ca2)"
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   332
    by (metis (no_types) `cs1 \<le> \<gamma>\<^isub>c ca1` join_ge1 le_post mono_rep_u order_trans post_map_acom)
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   333
  moreover have "post cs2 \<subseteq> \<gamma>\<^isub>u(post ca1 \<squnion> post ca2)"
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   334
    by (metis (no_types) `cs2 \<le> \<gamma>\<^isub>c ca2` join_ge2 le_post mono_rep_u order_trans post_map_acom)
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   335
  ultimately show ?case using `S \<subseteq> \<gamma>\<^isub>u sa` by (simp add: If.IH subset_iff)
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   336
next
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   337
  case (While b c1)
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   338
  then obtain cs1 ca1 I P Ia Pa where
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   339
    "cs = {I} WHILE b DO cs1 {P}" "ca = {Ia} WHILE b DO ca1 {Pa}"
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   340
    "I \<subseteq> \<gamma>\<^isub>u Ia" "P \<subseteq> \<gamma>\<^isub>u Pa" "cs1 \<le> \<gamma>\<^isub>c ca1"
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   341
    "strip cs1 = c1" "strip ca1 = c1"
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   342
    by (fastforce simp: strip_eq_While)
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   343
  moreover have "S \<union> post cs1 \<subseteq> \<gamma>\<^isub>u (sa \<squnion> post ca1)"
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   344
    using `S \<subseteq> \<gamma>\<^isub>u sa` le_post[OF `cs1 \<le> \<gamma>\<^isub>c ca1`, simplified]
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   345
    by (metis (no_types) join_ge1 join_ge2 le_sup_iff mono_rep_u order_trans)
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   346
  ultimately show ?case by (simp add: While.IH subset_iff)
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   347
qed
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   348
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   349
lemma step_preserves_le:
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   350
  "\<lbrakk> S \<subseteq> \<gamma>\<^isub>u sa; cs \<le> \<gamma>\<^isub>c ca; strip cs = c \<rbrakk>
45655
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
   351
   \<Longrightarrow> step S cs \<le> \<gamma>\<^isub>c(step' sa ca)"
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   352
by (metis le_strip step_preserves_le2 strip_acom)
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   353
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   354
lemma AI_sound: "AI c = Some c' \<Longrightarrow> CS UNIV c \<le> \<gamma>\<^isub>c c'"
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   355
proof(simp add: CS_def AI_def)
45655
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
   356
  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
   357
  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
   358
  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
   359
    by(simp add: strip_lpfpc[OF _ 1])
45655
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
   360
  have "lfp c (step UNIV) \<le> \<gamma>\<^isub>c (step' \<top> c')"
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   361
  proof(rule lfp_lowerbound[OF 3])
45655
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
   362
    show "step UNIV (\<gamma>\<^isub>c (step' \<top> c')) \<le> \<gamma>\<^isub>c (step' \<top> c')"
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   363
    proof(rule step_preserves_le[OF _ _ 3])
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   364
      show "UNIV \<subseteq> \<gamma>\<^isub>u \<top>" by simp
45655
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
   365
      show "\<gamma>\<^isub>c (step' \<top> c') \<le> \<gamma>\<^isub>c c'" by(rule mono_rep_c[OF 2])
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   366
    qed
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   367
  qed
45655
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
   368
  from this 2 show "lfp c (step UNIV) \<le> \<gamma>\<^isub>c c'"
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   369
    by (blast intro: mono_rep_c order_trans)
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   370
qed
45127
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
end
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   373
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   374
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   375
subsubsection "Monotonicity"
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   376
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   377
locale Abs_Int_Fun_mono = Abs_Int_Fun +
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   378
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
   379
begin
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   380
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   381
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
   382
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
   383
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   384
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
   385
by(simp add: le_fun_def)
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   386
45655
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
   387
lemma step'_mono: "S \<sqsubseteq> S' \<Longrightarrow> step' S c \<sqsubseteq> step' S' c"
45127
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   388
apply(induction c arbitrary: S S')
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45212
diff changeset
   389
apply (auto simp: Let_def mono_update mono_aval' le_join_disj split: option.split)
45127
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45113
diff changeset
   390
done
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   391
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   392
end
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   393
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   394
text{* Problem: not executable because of the comparison of abstract states,
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   395
i.e. functions, in the post-fixedpoint computation. *}
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   396
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   397
end