src/HOL/IMP/AbsInt1_ivl.thy
author nipkow
Fri, 02 Sep 2011 19:25:18 +0200
changeset 44656 22bbd0d1b943
child 44923 b80108b346a9
permissions -rw-r--r--
Added Abstract Interpretation theories
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
     1
(* Author: Tobias Nipkow *)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
     2
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
     3
theory AbsInt1_ivl
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
     4
imports AbsInt1
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
     5
begin
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
     6
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
     7
subsection "Interval Analysis"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
     8
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
     9
datatype ivl = I "int option" "int option"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    10
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    11
definition "rep_ivl i =
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    12
 (case i of I (Some l) (Some h) \<Rightarrow> {l..h} | I (Some l) None \<Rightarrow> {l..}
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    13
  | I None (Some h) \<Rightarrow> {..h} | I None None \<Rightarrow> UNIV)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    14
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    15
instantiation option :: (plus)plus
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    16
begin
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    17
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    18
fun plus_option where
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    19
"Some x + Some y = Some(x+y)" |
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    20
"_ + _ = None"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    21
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    22
instance proof qed
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    23
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    24
end
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    25
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    26
definition empty where "empty = I (Some 1) (Some 0)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    27
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    28
fun is_empty where
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    29
"is_empty(I (Some l) (Some h)) = (h<l)" |
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    30
"is_empty _ = False"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    31
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    32
lemma [simp]: "is_empty(I l h) =
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    33
  (case l of Some l \<Rightarrow> (case h of Some h \<Rightarrow> h<l | None \<Rightarrow> False) | None \<Rightarrow> False)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    34
by(auto split:option.split)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    35
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    36
lemma [simp]: "is_empty i \<Longrightarrow> rep_ivl i = {}"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    37
by(auto simp add: rep_ivl_def split: ivl.split option.split)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    38
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    39
definition "plus_ivl i1 i2 = (if is_empty i1 | is_empty i2 then empty
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    40
  else case (i1,i2) of (I l1 h1, I l2 h2) \<Rightarrow> I (l1+l2) (h1+h2))"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    41
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    42
instantiation ivl :: SL_top
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    43
begin
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    44
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    45
definition le_option :: "bool \<Rightarrow> int option \<Rightarrow> int option \<Rightarrow> bool" where
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    46
"le_option pos x y =
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    47
 (case x of (Some i) \<Rightarrow> (case y of (Some j) \<Rightarrow> (i<=j) | None \<Rightarrow> pos)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    48
  | None \<Rightarrow> (case y of Some j \<Rightarrow> (~pos) | None \<Rightarrow> True))"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    49
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    50
fun le_aux where
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    51
"le_aux (I l1 h1) (I l2 h2) = (le_option False l2 l1 & le_option True h1 h2)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    52
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    53
definition le_ivl where
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    54
"i1 \<sqsubseteq> i2 =
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    55
 (if is_empty i1 then True else
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    56
  if is_empty i2 then False else le_aux i1 i2)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    57
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    58
definition min_option :: "bool \<Rightarrow> int option \<Rightarrow> int option \<Rightarrow> int option" where
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    59
"min_option pos o1 o2 = (if le_option pos o1 o2 then o1 else o2)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    60
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    61
definition max_option :: "bool \<Rightarrow> int option \<Rightarrow> int option \<Rightarrow> int option" where
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    62
"max_option pos o1 o2 = (if le_option pos o1 o2 then o2 else o1)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    63
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    64
definition "i1 \<squnion> i2 =
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    65
 (if is_empty i1 then i2 else if is_empty i2 then i1
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    66
  else case (i1,i2) of (I l1 h1, I l2 h2) \<Rightarrow>
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    67
          I (min_option False l1 l2) (max_option True h1 h2))"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    68
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    69
definition "Top = I None None"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    70
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    71
instance
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    72
proof
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    73
  case goal1 thus ?case
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    74
    by(cases x, simp add: le_ivl_def le_option_def split: option.split)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    75
next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    76
  case goal2 thus ?case
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    77
    by(cases x, cases y, cases z, auto simp: le_ivl_def le_option_def split: option.splits if_splits)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    78
next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    79
  case goal3 thus ?case
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    80
    by(cases x, cases y, simp add: le_ivl_def join_ivl_def le_option_def min_option_def max_option_def split: option.splits)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    81
next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    82
  case goal4 thus ?case
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    83
    by(cases x, cases y, simp add: le_ivl_def join_ivl_def le_option_def min_option_def max_option_def split: option.splits)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    84
next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    85
  case goal5 thus ?case
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    86
    by(cases x, cases y, cases z, auto simp add: le_ivl_def join_ivl_def le_option_def min_option_def max_option_def split: option.splits if_splits)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    87
next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    88
  case goal6 thus ?case
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    89
    by(cases x, simp add: Top_ivl_def le_ivl_def le_option_def split: option.split)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    90
qed
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    91
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    92
end
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    93
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    94
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    95
instantiation ivl :: L_top_bot
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    96
begin
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    97
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    98
definition "i1 \<sqinter> i2 = (if is_empty i1 | is_empty i2 then empty
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    99
  else case (i1,i2) of (I l1 h1, I l2 h2) \<Rightarrow> I (max_option False l1 l2) (min_option True h1 h2))"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   100
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   101
definition "Bot = empty"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   102
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   103
instance
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   104
proof
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   105
  case goal1 thus ?case
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   106
    by (simp add:meet_ivl_def empty_def meet_ivl_def le_ivl_def le_option_def max_option_def min_option_def split: ivl.splits option.splits)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   107
next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   108
  case goal2 thus ?case
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   109
    by (simp add:meet_ivl_def empty_def meet_ivl_def le_ivl_def le_option_def max_option_def min_option_def split: ivl.splits option.splits)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   110
next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   111
  case goal3 thus ?case
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   112
    by (cases x, cases y, cases z, auto simp add: le_ivl_def meet_ivl_def empty_def le_option_def max_option_def min_option_def split: option.splits if_splits)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   113
next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   114
  case goal4 show ?case by(cases x, simp add: Bot_ivl_def empty_def le_ivl_def)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   115
qed
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   116
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   117
end
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   118
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   119
instantiation option :: (minus)minus
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   120
begin
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   121
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   122
fun minus_option where
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   123
"Some x - Some y = Some(x-y)" |
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   124
"_ - _ = None"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   125
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   126
instance proof qed
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   127
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   128
end
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   129
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   130
definition "minus_ivl i1 i2 =
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   131
 (if is_empty i1 | is_empty i2 then empty
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   132
  else case (i1,i2) of
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   133
    (I l1 h1, I l2 h2) \<Rightarrow> I (l1-h2) (h1-l2))"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   134
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   135
lemma rep_minus_ivl:
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   136
  "n1 : rep_ivl i1 \<Longrightarrow> n2 : rep_ivl i2 \<Longrightarrow> n1-n2 : rep_ivl(minus_ivl i1 i2)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   137
by(auto simp add: minus_ivl_def rep_ivl_def split: ivl.splits option.splits)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   138
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   139
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   140
definition "inv_plus_ivl i1 i2 i =
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   141
  (if is_empty i then empty
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   142
   else i1 \<sqinter> minus_ivl i i2, i2 \<sqinter> minus_ivl i i1)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   143
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   144
fun inv_less_ivl :: "ivl \<Rightarrow> ivl \<Rightarrow> bool \<Rightarrow> ivl * ivl" where
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   145
"inv_less_ivl (I l1 h1) (I l2 h2) res =
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   146
  (if res
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   147
   then (I l1 (min_option True h1 (h2 - Some 1)), I (max_option False (l1 + Some 1) l2) h2)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   148
   else (I (max_option False l1 l2) h1, I l2 (min_option True h1 h2)))"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   149
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   150
interpretation Rep rep_ivl
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   151
proof
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   152
  case goal1 thus ?case
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   153
    by(auto simp: rep_ivl_def le_ivl_def le_option_def split: ivl.split option.split if_splits)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   154
qed
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   155
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   156
interpretation Val_abs rep_ivl "%n. I (Some n) (Some n)" plus_ivl
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   157
proof
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   158
  case goal1 thus ?case by(simp add: rep_ivl_def)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   159
next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   160
  case goal2 thus ?case
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   161
    by(auto simp add: rep_ivl_def plus_ivl_def split: ivl.split option.splits)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   162
qed
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   163
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   164
interpretation Rep1 rep_ivl
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   165
proof
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   166
  case goal1 thus ?case
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   167
    by(auto simp add: rep_ivl_def meet_ivl_def empty_def min_option_def max_option_def split: ivl.split option.split)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   168
next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   169
  case goal2 show ?case by(auto simp add: Bot_ivl_def rep_ivl_def empty_def)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   170
qed
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   171
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   172
interpretation Val_abs1 rep_ivl "%n. I (Some n) (Some n)" plus_ivl inv_plus_ivl inv_less_ivl
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   173
proof
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   174
  case goal1 thus ?case
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   175
    by(auto simp add: inv_plus_ivl_def)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   176
      (metis rep_minus_ivl add_diff_cancel add_commute)+
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   177
next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   178
  case goal2 thus ?case
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   179
    by(cases a1, cases a2,
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   180
      auto simp: rep_ivl_def min_option_def max_option_def le_option_def split: if_splits option.splits)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   181
qed
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   182
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   183
interpretation Abs_Int1 rep_ivl "%n. I (Some n) (Some n)" plus_ivl inv_plus_ivl inv_less_ivl
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   184
defines afilter_ivl is afilter
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   185
and bfilter_ivl is bfilter
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   186
and AI_ivl is AI
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   187
and aval_ivl is aval'
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   188
..
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   189
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   190
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   191
fun list_up where
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   192
"list_up bot = bot" |
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   193
"list_up (Up x) = Up(list x)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   194
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   195
value [code] "list_up(afilter_ivl (N 5) (I (Some 4) (Some 5)) Top)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   196
value [code] "list_up(afilter_ivl (N 5) (I (Some 4) (Some 4)) Top)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   197
value [code] "list_up(afilter_ivl (V ''x'') (I (Some 4) (Some 4))
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   198
 (Up(FunDom(Top(''x'':=I (Some 5) (Some 6))) [''x''])))"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   199
value [code] "list_up(afilter_ivl (V ''x'') (I (Some 4) (Some 5))
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   200
 (Up(FunDom(Top(''x'':=I (Some 5) (Some 6))) [''x''])))"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   201
value [code] "list_up(afilter_ivl (Plus (V ''x'') (V ''x'')) (I (Some 0) (Some 10))
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   202
  (Up(FunDom(Top(''x'':= I (Some 0) (Some 100)))[''x''])))"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   203
value [code] "list_up(afilter_ivl (Plus (V ''x'') (N 7)) (I (Some 0) (Some 10))
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   204
  (Up(FunDom(Top(''x'':= I (Some 0) (Some 100)))[''x''])))"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   205
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   206
value [code] "list_up(bfilter_ivl (Less (V ''x'') (V ''x'')) True
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   207
  (Up(FunDom(Top(''x'':= I (Some 0) (Some 0)))[''x''])))"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   208
value [code] "list_up(bfilter_ivl (Less (V ''x'') (V ''x'')) True
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   209
  (Up(FunDom(Top(''x'':= I (Some 0) (Some 2)))[''x''])))"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   210
value [code] "list_up(bfilter_ivl (Less (V ''x'') (Plus (N 10) (V ''y''))) True
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   211
  (Up(FunDom(Top(''x'':= I (Some 15) (Some 20),''y'':= I (Some 5) (Some 7)))[''x'', ''y''])))"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   212
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   213
definition "test_ivl1 =
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   214
 ''y'' ::= N 7;
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   215
 IF Less (V ''x'') (V ''y'')
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   216
 THEN ''y'' ::= Plus (V ''y'') (V ''x'')
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   217
 ELSE ''x'' ::= Plus (V ''x'') (V ''y'')"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   218
value [code] "list_up(AI_ivl test_ivl1 Top)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   219
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   220
value "list_up (AI_ivl test3_const Top)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   221
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   222
value "list_up (AI_ivl test5_const Top)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   223
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   224
value "list_up (AI_ivl test6_const Top)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   225
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   226
definition "test2_ivl =
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   227
 ''y'' ::= N 7;
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   228
 WHILE Less (V ''x'') (N 100) DO ''y'' ::= Plus (V ''y'') (N 1)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   229
value [code] "list_up(AI_ivl test2_ivl Top)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   230
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   231
end